From 0b2ba4d4944014e0b4b04e26f180e18b45677e15 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 23 Nov 2023 18:55:40 +0000 Subject: [PATCH 001/189] feat: make cones irreducible --- CvxLean/Lib/Cones/ExpCone.lean | 4 +++- CvxLean/Lib/Cones/PSDCone.lean | 5 +++-- CvxLean/Lib/Cones/PosOrthCone.lean | 3 +++ CvxLean/Lib/Cones/SOCone.lean | 18 +++++++++++++----- CvxLean/Lib/Cones/ZeroCone.lean | 13 ++++++++++--- 5 files changed, 32 insertions(+), 11 deletions(-) diff --git a/CvxLean/Lib/Cones/ExpCone.lean b/CvxLean/Lib/Cones/ExpCone.lean index d7cc10d6..fc7e3551 100644 --- a/CvxLean/Lib/Cones/ExpCone.lean +++ b/CvxLean/Lib/Cones/ExpCone.lean @@ -2,13 +2,15 @@ import Mathlib.Data.Complex.Exponential namespace Real -@[reducible] +@[irreducible] def expCone (x y z : Real) : Prop := (0 < y ∧ y * exp (x / y) ≤ z) ∨ (y = 0 ∧ 0 ≤ z ∧ x ≤ 0) +@[irreducible] def Vec.expCone (x y z : Fin n → Real) : Prop := ∀ i, Real.expCone (x i) (y i) (z i) +@[irreducible] theorem exp_iff_expCone (t x : Real) : exp x ≤ t ↔ expCone x 1 t := by unfold expCone rw [iff_def] diff --git a/CvxLean/Lib/Cones/PSDCone.lean b/CvxLean/Lib/Cones/PSDCone.lean index 3cc870f6..294a78d4 100644 --- a/CvxLean/Lib/Cones/PSDCone.lean +++ b/CvxLean/Lib/Cones/PSDCone.lean @@ -1,9 +1,10 @@ import Mathlib.Data.Real.Basic import Mathlib.LinearAlgebra.Matrix.PosDef -namespace Real +namespace Real -def Matrix.PSDCone {m} [Fintype m] (M : Matrix m m Real) : Prop := +@[irreducible] +def Matrix.PSDCone {m} [Fintype m] (M : Matrix m m Real) : Prop := Matrix.PosSemidef M end Real diff --git a/CvxLean/Lib/Cones/PosOrthCone.lean b/CvxLean/Lib/Cones/PosOrthCone.lean index 4e28667e..730ff49b 100644 --- a/CvxLean/Lib/Cones/PosOrthCone.lean +++ b/CvxLean/Lib/Cones/PosOrthCone.lean @@ -3,12 +3,15 @@ import CvxLean.Lib.Math.Data.Matrix namespace Real +@[irreducible] def posOrthCone (x : Real) : Prop := 0 ≤ x +@[irreducible] def Vec.posOrthCone (x : Fin n → Real) : Prop := ∀ i, Real.posOrthCone (x i) +@[irreducible] def Matrix.posOrthCone (M : Matrix (Fin m) (Fin n) Real) : Prop := ∀ i j, Real.posOrthCone (M i j) diff --git a/CvxLean/Lib/Cones/SOCone.lean b/CvxLean/Lib/Cones/SOCone.lean index 7bf543cf..ef26e0ce 100644 --- a/CvxLean/Lib/Cones/SOCone.lean +++ b/CvxLean/Lib/Cones/SOCone.lean @@ -14,12 +14,24 @@ open BigOperators variable [Fintype m] [Fintype n] +@[irreducible] def soCone (t : Real) (x : n → Real) : Prop := sqrt (∑ i, x i ^ 2) ≤ t +@[irreducible] def rotatedSoCone (v w : Real) (x : n → Real) : Prop := (∑ i, x i ^ 2) ≤ (v * w) * 2 ∧ 0 ≤ v ∧ 0 ≤ w +@[irreducible] +def Vec.soCone (t : m → Real) (X : Matrix m n Real) : Prop := + ∀ i, Real.soCone (t i) (X i) + +@[irreducible] +def Vec.rotatedSoCone (v w : m → Real) (X : Matrix m n Real) : Prop := + ∀ i, Real.rotatedSoCone (v i) (w i) (X i) + +section ConeConversion + noncomputable def rotateSoCone {n : ℕ} (t : Real) (x : Fin n.succ → Real) : Real × Real × (Fin n → Real) := ((t + x 0) / sqrt 2, (t - x 0) / sqrt 2, fun i => x i.succ) @@ -28,11 +40,7 @@ noncomputable def unrotateSoCone {n : ℕ} (v w : Real) (x : Fin n → Real) : Real × (Fin n.succ → Real) := ((v + w) / sqrt 2, Matrix.vecCons ((v - w) / sqrt 2) x) -def Vec.soCone (t : m → Real) (X : Matrix m n Real) : Prop := - ∀ i, Real.soCone (t i) (X i) - -def Vec.rotatedSoCone (v w : m → Real) (X : Matrix m n Real) : Prop := - ∀ i, Real.rotatedSoCone (v i) (w i) (X i) +end ConeConversion section Lemmas diff --git a/CvxLean/Lib/Cones/ZeroCone.lean b/CvxLean/Lib/Cones/ZeroCone.lean index 48491226..1fc22dca 100644 --- a/CvxLean/Lib/Cones/ZeroCone.lean +++ b/CvxLean/Lib/Cones/ZeroCone.lean @@ -1,11 +1,18 @@ import Mathlib.Data.Real.Basic +import CvxLean.Lib.Math.Data.Matrix -namespace Real +namespace Real +@[irreducible] def zeroCone (x : Real) : Prop := x = 0 -def Vec.zeroCone (x : Fin n → Real) : Prop := +@[irreducible] +def Vec.zeroCone (x : Fin n → Real) : Prop := ∀ i, Real.zeroCone (x i) -end Real +@[irreducible] +def Matrix.zeroCone (M : Matrix (Fin n) (Fin m) Real) : Prop := + ∀ i j, Real.zeroCone (M i j) + +end Real From 5edb3ac39f20cf294fe02333e238fb7e3a953127 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 23 Nov 2023 18:56:10 +0000 Subject: [PATCH 002/189] wip: ensure cones are not unfolded --- CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean | 1 + CvxLean/Tactic/DCP/AtomLibrary/Sets/Le.lean | 13 ++++++++----- CvxLean/Tactic/DCP/Atoms.lean | 6 ++++++ CvxLean/Test/PreDCP/DQCP.lean | 2 ++ 4 files changed, 17 insertions(+), 5 deletions(-) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean index 366327e2..95ff4b14 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean @@ -55,6 +55,7 @@ vconditionElimination -- NOTE(RFM): It is not straightforward to express it in terms of x ^ (-1) and -- x ^ 2 because of vconditions. +set_option trace.Meta.debug true in declare_atom powNegTwo [convex] (x : ℝ)- : x ^ (-2) := vconditions (hx : 0 < x) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Sets/Le.lean b/CvxLean/Tactic/DCP/AtomLibrary/Sets/Le.lean index 77eb39fd..e5dbaa12 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Sets/Le.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Sets/Le.lean @@ -4,6 +4,7 @@ import CvxLean.Tactic.DCP.AtomLibrary.Fns.Sub namespace CvxLean +set_option trace.Meta.debug true in declare_atom le [convex_set] (x : ℝ)- (y : ℝ)+ : x ≤ y := vconditions implementationVars @@ -11,11 +12,13 @@ implementationObjective Real.posOrthCone (y - x) implementationConstraints solution solutionEqualsAtom by - simp [Real.posOrthCone] + unfold Real.posOrthCone + simp feasibility optimality by intros x' y' hx hy h - simp [Real.posOrthCone] at h + unfold Real.posOrthCone at h + simp at h exact (hx.trans h).trans hy vconditionElimination @@ -32,10 +35,10 @@ solutionEqualsAtom by constructor · intros h i rw [← le.solEqAtom] + unfold Real.posOrthCone apply h - · intros h i - erw [le.solEqAtom] - apply h + · intros hy i + simp [hy i] feasibility optimality by intros x' y' hx hy h i diff --git a/CvxLean/Tactic/DCP/Atoms.lean b/CvxLean/Tactic/DCP/Atoms.lean index 69ddd983..6cf348d6 100644 --- a/CvxLean/Tactic/DCP/Atoms.lean +++ b/CvxLean/Tactic/DCP/Atoms.lean @@ -159,6 +159,9 @@ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandEla mkLambdaFVars xs <| ← mkLambdaFVars cs adjustedFeas oldFeasibilityAdjusted := oldFeasibilityAdjusted.push adjustedFeas + for proof in solEqAtomProofs do + trace[Meta.debug] "solEqAtomProofs: {← inferType proof}" + for feas in oldFeasibilityAdjusted do trace[Meta.debug] "oldFeasibilityAdjusted: {← inferType feas}" @@ -354,6 +357,7 @@ def elabSolEqAtom (argDecls : Array LocalDecl) (vconds : Array (Lean.Name × Exp let sols := sols.map (mkAppNBeta · xs) let impObj' := convertLambdasToLets impObj sols let ty ← mkEq impObj' body + trace[Meta.debug] "ensuring type {ty}" let solEqAtom ← Elab.Term.elabTermAndSynthesizeEnsuringType stx (some ty) return ← mkLambdaFVars xs $ ← mkLambdaFVars cs solEqAtom @@ -527,6 +531,8 @@ def elabVCondElim (curv : Curvature) (argDecls : Array LocalDecl) (vconds : Arra -- | Curvature.ConvexSet => Curvature.ConvexSet -- | _ => Curvature.Affine + trace[Meta.debug] "before reduce sol eq atom: {atomData.solEqAtom}" + let atomData ← reduceAtomData objCurv atomData -- trace[Meta.debug] "HERE Reduced atom: {atomData}" let atomData ← addAtomDataDecls id.getId atomData diff --git a/CvxLean/Test/PreDCP/DQCP.lean b/CvxLean/Test/PreDCP/DQCP.lean index 4a6ac6a8..4224b07f 100644 --- a/CvxLean/Test/PreDCP/DQCP.lean +++ b/CvxLean/Test/PreDCP/DQCP.lean @@ -133,6 +133,8 @@ time_cmd reduction redqcp2/dqcp2 : hypersonicShapeDesign 0.35 0.65 := by convexify dcp --TODO: This has 0 ≤ Δx ????? +-- solve dqcp2 + end QCP2 end From 7d0a944858f6311558e2cc75fb2176ec5de2285f Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 23 Nov 2023 19:00:08 +0000 Subject: [PATCH 003/189] fix: sets with irreducible cones --- CvxLean/Lib/Cones/SOCone.lean | 2 +- .../Tactic/DCP/AtomLibrary/Sets/Cones.lean | 20 +++++++------------ CvxLean/Tactic/DCP/AtomLibrary/Sets/Eq.lean | 8 +++++--- CvxLean/Tactic/DCP/AtomLibrary/Sets/Le.lean | 4 ++-- 4 files changed, 15 insertions(+), 19 deletions(-) diff --git a/CvxLean/Lib/Cones/SOCone.lean b/CvxLean/Lib/Cones/SOCone.lean index ef26e0ce..ad6ace6c 100644 --- a/CvxLean/Lib/Cones/SOCone.lean +++ b/CvxLean/Lib/Cones/SOCone.lean @@ -50,7 +50,7 @@ which is equivalent to z ^ 2 ≤ x * y. -/ lemma soCone_add_sub_two_mul_of_nonneg {x y : ℝ} (z : ℝ) (hx : 0 ≤ x) (hy : 0 ≤ y) : soCone (x + y) ![x - y, 2 * z] ↔ z ^ (2 : ℝ) ≤ x * y := by have hxy := add_nonneg hx hy - conv => lhs; simp [soCone, sqrt_le_left hxy, ←le_sub_iff_add_le'] + conv => lhs; unfold soCone; simp [sqrt_le_left hxy, ←le_sub_iff_add_le'] ring_nf; simp lemma soCone_add_sub_two_of_nonneg {x y : ℝ} (hx : 0 ≤ x) (hy : 0 ≤ y) : diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Sets/Cones.lean b/CvxLean/Tactic/DCP/AtomLibrary/Sets/Cones.lean index 157db48b..9441c646 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Sets/Cones.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Sets/Cones.lean @@ -39,13 +39,11 @@ end PosOrthCone section ExpCone declare_atom expCone [cone] (x : ℝ)? (y : ℝ)? (z : ℝ)? : expCone x y z := -optimality by - simp [expCone] +optimality le_refl _ declare_atom Vec.expCone [cone] (n : Nat)& (x : (Fin n) → ℝ)? (y : (Fin n) → ℝ)? (z : (Fin n) → ℝ)? : Vec.expCone x y z := -optimality by - simp [Vec.expCone] +optimality le_refl _ end ExpCone @@ -54,23 +52,19 @@ section SOCone declare_atom soCone [cone] (n : Nat)& (t : ℝ)? (x : (Fin n) → ℝ)? : soCone t x := -optimality by - simp [soCone] +optimality le_refl _ declare_atom Vec.soCone [cone] (n : Nat)& (m : Nat)& (t : Fin m → Real)? (X : Matrix.{0,0,0} (Fin m) (Fin n) Real)? : Vec.soCone t X := -optimality by - simp [Vec.soCone] +optimality le_refl _ declare_atom rotatedSoCone [cone] (n : Nat)& (v : ℝ)? (w : ℝ)? (x : (Fin n) → ℝ)? : rotatedSoCone v w x := -optimality by - simp [rotatedSoCone] +optimality le_refl _ declare_atom Vec.rotatedSoCone [cone] (m : Nat)& (n : Nat)& (v : (Fin n) → ℝ)? (w : (Fin n) → ℝ)? (x : (Fin n) → (Fin m) → ℝ)? : Vec.rotatedSoCone v w x := -optimality by - simp [Vec.rotatedSoCone] +optimality le_refl _ end SOCone @@ -79,7 +73,7 @@ section PSDCone declare_atom Matrix.PSDCone [cone] (m : Type)& (hm : Fintype.{0} m)& (A : Matrix.{0,0,0} m m ℝ)? : Matrix.PSDCone A := -optimality fun h => h +optimality le_refl _ end PSDCone diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Sets/Eq.lean b/CvxLean/Tactic/DCP/AtomLibrary/Sets/Eq.lean index 5b8b06e8..737d605f 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Sets/Eq.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Sets/Eq.lean @@ -11,11 +11,13 @@ implementationObjective Real.zeroCone (y - x) implementationConstraints solution solutionEqualsAtom by - simp [Real.zeroCone, sub_eq_iff_eq_add, zero_add] - exact Iff.intro Eq.symm Eq.symm; + unfold Real.zeroCone + simp [sub_eq_iff_eq_add, zero_add] + exact Iff.intro Eq.symm Eq.symm feasibility optimality by - simp [Real.zeroCone, sub_eq_iff_eq_add, zero_add] + unfold Real.zeroCone + simp [sub_eq_iff_eq_add, zero_add] exact Eq.symm vconditionElimination diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Sets/Le.lean b/CvxLean/Tactic/DCP/AtomLibrary/Sets/Le.lean index e5dbaa12..3f23d67b 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Sets/Le.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Sets/Le.lean @@ -4,7 +4,6 @@ import CvxLean.Tactic.DCP.AtomLibrary.Fns.Sub namespace CvxLean -set_option trace.Meta.debug true in declare_atom le [convex_set] (x : ℝ)- (y : ℝ)+ : x ≤ y := vconditions implementationVars @@ -42,7 +41,8 @@ solutionEqualsAtom by feasibility optimality by intros x' y' hx hy h i - apply le.optimality _ _ _ _ (hx i) (hy i) (h i) + unfold Real.Vec.posOrthCone at h + exact le.optimality _ _ _ _ (hx i) (hy i) (h i) vconditionElimination end CvxLean From 61e3cb7cf0dfe93785ce80084318fe30bc7c7cbd Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 23 Nov 2023 19:46:24 +0000 Subject: [PATCH 004/189] fix: atom declarations with irreducible cones --- CvxLean/Tactic/DCP/AtomLibrary/Fns/Abs.lean | 11 +++- CvxLean/Tactic/DCP/AtomLibrary/Fns/Entr.lean | 15 ++--- CvxLean/Tactic/DCP/AtomLibrary/Fns/Exp.lean | 6 +- CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean | 34 ++++++++--- CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean | 59 +++++++++++++++---- CvxLean/Tactic/DCP/AtomLibrary/Fns/Log.lean | 15 +++-- .../Tactic/DCP/AtomLibrary/Fns/LogDet.lean | 5 +- CvxLean/Tactic/DCP/AtomLibrary/Fns/Norm.lean | 6 +- .../DCP/AtomLibrary/Fns/PosSemidef.lean | 4 +- CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean | 6 +- CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean | 9 ++- CvxLean/Tactic/DCP/AtomLibrary/Fns/Sqrt.lean | 14 +++-- CvxLean/Tactic/DCP/AtomLibrary/Fns/XExp.lean | 9 ++- CvxLean/Tactic/DCP/Atoms.lean | 4 +- 14 files changed, 139 insertions(+), 58 deletions(-) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Abs.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Abs.lean index b7d36d5b..f5736ac2 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Abs.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Abs.lean @@ -29,6 +29,7 @@ feasibility optimality by apply abs_le.2 rw [←sub_nonneg, sub_neg_eq_add, add_comm, ←sub_nonneg (b := x)] + unfold posOrthCone at c_pos c_neg exact ⟨c_neg, c_pos⟩ vconditionElimination @@ -43,13 +44,16 @@ solution (t := abs x) solutionEqualsAtom rfl feasibility (c_pos : by - intros _ _ + unfold Real.Vec.posOrthCone + intros apply abs.feasibility0) (c_neg : by - intros _ _ + unfold Real.Vec.posOrthCone + intros apply abs.feasibility1) optimality by intros i + unfold Real.Vec.posOrthCone at c_pos c_neg apply abs.optimality _ _ (c_pos i) (c_neg i) vconditionElimination @@ -66,13 +70,16 @@ solution (T := M.abs) solutionEqualsAtom rfl feasibility (c_pos : by + unfold Real.Matrix.posOrthCone intros _ _ _ apply abs.feasibility0) (c_neg : by + unfold Real.Matrix.posOrthCone intros _ _ _ apply abs.feasibility1) optimality by intros i j + unfold Real.Matrix.posOrthCone at c_pos c_neg apply abs.optimality _ _ (c_pos i j) (c_neg i j) vconditionElimination diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Entr.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Entr.lean index ddd1eddc..57a28ed0 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Entr.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Entr.lean @@ -17,7 +17,7 @@ solution (t := -(x * log x)) solutionEqualsAtom by rfl feasibility (c : by - simp [expCone] + unfold expCone cases (lt_or_eq_of_le cond) with | inl h => left @@ -31,7 +31,7 @@ feasibility refine ⟨h, ?_⟩ simp [h]) optimality by - simp [expCone] at c + unfold expCone at c cases c with | inl c => rw [mul_comm, ←neg_mul, ←div_le_iff c.1] @@ -42,7 +42,7 @@ optimality by simp [entr, c.1, c.2] vconditionElimination (cond : by - simp [expCone] at c + unfold expCone at c cases c with | inl c => exact le_of_lt c.1 | inr c => exact le_of_eq c.1.symm) @@ -71,16 +71,17 @@ solution (t := Vec.entr x) solutionEqualsAtom by rfl feasibility (c : by - simp [Vec.expCone, Vec.entr] - intros i + unfold Vec.expCone Vec.entr + intros t i exact entr.feasibility0 (x i) (cond i)) optimality by - simp [Vec.expCone, Vec.entr] at * + unfold Vec.expCone at c + unfold Vec.entr intros i exact entr.optimality (x i) (t i) (c i) vconditionElimination (cond : by - simp [Vec.expCone, Vec.entr] at c + unfold Vec.expCone at c intros i exact entr.vcondElim0 (x i) (t i) (c i)) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Exp.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Exp.lean index c27cf197..2e15d69e 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Exp.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Exp.lean @@ -6,7 +6,6 @@ namespace CvxLean open Real -set_option trace.Meta.debug true declare_atom exp [convex] (x : ℝ)+ : Real.exp x := vconditions implementationVars (t : ℝ) @@ -16,7 +15,8 @@ solution (t := exp x) solutionEqualsAtom by rfl; feasibility - (c_exp : by simp [expCone]) + (c_exp : by + unfold expCone; simp) optimality by intros x' hx rw [←exp_iff_expCone] at c_exp @@ -36,10 +36,12 @@ solutionEqualsAtom rfl feasibility (c_exp: by + unfold Vec.exp Vec.expCone intros _ i apply exp.feasibility0) optimality by intros x' hx i + unfold Vec.expCone at c_exp apply exp.optimality _ _ (c_exp i) _ (hx i) vconditionElimination diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean index 7e1c761c..ec9f5bce 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean @@ -81,23 +81,39 @@ solutionEqualsAtom by simp [Vec.huber, ←huber.solEqAtom (x i)] feasibility (c1 : by - simpa using (fun i => huber.feasibility0 (x i))) + dsimp + intros i + have h := huber.feasibility0 (x i) + unfold Real.posOrthCone at h + simpa using h) (c2 : by - simpa using (fun i => huber.feasibility1 (x i))) + dsimp + intros i + have h := huber.feasibility1 (x i) + unfold Real.posOrthCone at h + simpa using h) (c3 : by - simpa using (fun i => huber.feasibility2 (x i))) + dsimp + intros i + have h := huber.feasibility2 (x i) + unfold Real.posOrthCone at h + simpa using h) (c4 : by - simpa using (fun i => huber.feasibility3 (x i))) + dsimp + intros i + have h := huber.feasibility3 (x i) + unfold Real.posOrthCone at h + simpa using h) optimality by intros i simp [Vec.huber] rw [←rpow_two] apply huber.optimality (x i) (v i) (w i) ((w i) ^ 2) - { simpa [posOrthCone] using c1 i } - { simpa [posOrthCone] using c2 i } - { simpa [posOrthCone] using c3 i } - { simpa [posOrthCone] using c4 i } - { simp [rotatedSoCone]; norm_num [sq_nonneg] } + { unfold posOrthCone; simpa using c1 i } + { unfold posOrthCone; simpa using c2 i } + { unfold posOrthCone; simpa using c3 i } + { unfold posOrthCone; simpa using c4 i } + { unfold rotatedSoCone; simp [sq_nonneg]; norm_num } vconditionElimination end CvxLean diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean index 32dfe611..7b292fd9 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean @@ -10,7 +10,6 @@ namespace CvxLean open Real -set_option trace.Meta.debug true in declare_atom klDiv [convex] (x : ℝ)? (y : ℝ)? : x * log (x / y) - x + y := vconditions (hx : 0 ≤ x) @@ -26,6 +25,8 @@ solutionEqualsAtom by ring feasibility (c1 : by + dsimp + unfold expCone cases (lt_or_eq_of_le hx) with | inl hx => left @@ -40,6 +41,7 @@ feasibility (c2 : by simp [exp_log hy]) optimality by + unfold expCone at c1 cases c1 with | inl c => have hxpos := c.1 @@ -58,7 +60,7 @@ optimality by simp [c.1, c.2.2] vconditionElimination (hx : by - simp [expCone] at c1 + unfold expCone at c1 cases c1 with | inl c => exact le_of_lt c.1 | inr c => rw [c.1]) @@ -80,14 +82,30 @@ solutionEqualsAtom by ring feasibility (c1 : klDiv.feasibility0 x y hx hy) - (c2 : by simpa [*] using klDiv.feasibility1 x y hx hy) + (c2 : by + dsimp + have h := klDiv.feasibility1 x y hx hy + unfold posOrthCone at h + simpa using h) optimality by - apply klDiv.optimality x y t y' (exp y') c1 <;> simpa [expCone, posOrthCone] + apply klDiv.optimality x y t y' (exp y') c1 + { unfold posOrthCone expCone at * + simpa using c2 } + { unfold expCone at * + simp } vconditionElimination (hx : by - apply klDiv.vcondElim0 x y t y' (exp y') c1 <;> simpa [expCone, posOrthCone]) + apply klDiv.vcondElim0 x y t y' (exp y') c1 + { unfold posOrthCone expCone at * + simpa using c2 } + { unfold expCone at * + simp }) (hy : by - apply klDiv.vcondElim1 x y t y' (exp y') c1 <;> simpa [expCone, posOrthCone]) + apply klDiv.vcondElim1 x y t y' (exp y') c1 + { unfold posOrthCone expCone at * + simpa using c2 } + { unfold expCone at * + simp }) declare_atom Vec.klDiv [convex] (m : Nat)& (x : Fin m → ℝ)? (y : Fin m → ℝ)? : Vec.klDiv x y := @@ -105,21 +123,36 @@ solutionEqualsAtom by feasibility (c1 : by simp [Vec.klDiv, klDiv] + unfold Vec.expCone intros i exact (klDiv.feasibility0 (x i) (y i) (hx i) (hy i))) (c2 : by simp [Vec.klDiv, klDiv] intros i - simpa [*] using klDiv.feasibility1 (x i) (y i) (hx i) (hy i)) + have h := klDiv.feasibility1 (x i) (y i) (hx i) (hy i) + unfold posOrthCone at h + simpa using h) optimality fun i => by - apply klDiv.optimality (x i) (y i) (t i) (y' i) (exp (y' i)) (c1 i) <;> - simpa [posOrthCone, expCone] using c2 i + unfold Vec.expCone at c1 + apply klDiv.optimality (x i) (y i) (t i) (y' i) (exp (y' i)) (c1 i) + { unfold posOrthCone expCone at * + simpa using (c2 i) } + { unfold expCone at * + simp } vconditionElimination (hx : fun i => by - apply klDiv.vcondElim0 (x i) (y i) (t i) (y' i) (exp (y' i)) (c1 i) <;> - simpa [posOrthCone, expCone] using (c2 i)) + unfold Vec.expCone at c1 + apply klDiv.vcondElim0 (x i) (y i) (t i) (y' i) (exp (y' i)) (c1 i) + { unfold posOrthCone expCone at * + simpa using (c2 i) } + { unfold expCone at * + simp }) (hy : fun i => by - apply klDiv.vcondElim1 (x i) (y i) (t i) (y' i) (exp (y' i)) (c1 i) <;> - simpa [posOrthCone, expCone] using (c2 i)) + unfold Vec.expCone at c1 + apply klDiv.vcondElim1 (x i) (y i) (t i) (y' i) (exp (y' i)) (c1 i) + { unfold posOrthCone expCone at * + simpa using (c2 i) } + { unfold expCone at * + simp }) end CvxLean diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Log.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Log.lean index 2e98b252..05acde8e 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Log.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Log.lean @@ -15,11 +15,13 @@ solution (t := log x) solutionEqualsAtom by rfl; feasibility (c_exp : by - simp [expCone] + unfold expCone + simp rw [Real.exp_log cond]) optimality by intros y hy - simp [expCone] at c_exp + unfold expCone at c_exp + simp at c_exp have hxpos := lt_of_lt_of_le (exp_pos t) c_exp have hypos := lt_of_lt_of_le hxpos hy have hexptley := le_trans c_exp hy @@ -27,7 +29,8 @@ optimality by vconditionElimination (cond : by intros y hy - simp [expCone] at c_exp + unfold expCone at c_exp + simp at c_exp have hexppos := exp_pos t have hxpos := lt_of_lt_of_le hexppos c_exp exact lt_of_lt_of_le hxpos hy) @@ -43,14 +46,18 @@ solution (t := log x) solutionEqualsAtom rfl feasibility (c_exp: by - intros _ i + dsimp + unfold Vec.expCone + intros i apply log.feasibility0 apply cond) optimality by intros x' hx i + unfold Vec.expCone at c_exp apply log.optimality _ _ (c_exp i) _ (hx i) vconditionElimination (cond : by intros x' hx i + unfold Vec.expCone at c_exp apply log.vcondElim0 _ _ (c_exp i) _ (hx i)) end CvxLean diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogDet.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogDet.lean index ffbd8086..ed6a46f9 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogDet.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogDet.lean @@ -41,7 +41,8 @@ solutionEqualsAtom by congr feasibility (c_exp : by - simp only [Real.Vec.expCone, dif_pos hA] + unfold Real.Vec.expCone + simp only [dif_pos hA] intro i show Real.expCone ((Real.log (LDL.diagEntries hA i))) 1 (Matrix.diag ((LDL.diag hA) * (LDL.lower hA)ᵀ) i) @@ -55,6 +56,7 @@ optimality by have ht : ∀ (i : Fin n), Real.exp (t i) ≤ Matrix.diag Y i := by intro i rw [Real.exp_iff_expCone] + unfold Real.Vec.expCone at c_exp apply c_exp dsimp at c_posdef -- NOTE: Not matching instances, hence `convert` not `exact`. @@ -66,6 +68,7 @@ vconditionElimination have ht : ∀ (i : Fin n), Real.exp (t i) ≤ Matrix.diag Y i := by intro i rw [Real.exp_iff_expCone] + unfold Real.Vec.expCone at c_exp apply c_exp exact Matrix.LogDetAtom.cond_elim (A := A) (t := t) (Y := Y) ht rfl rfl (by convert c_posdef)) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Norm.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Norm.lean index 9d842866..b41df638 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Norm.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Norm.lean @@ -16,9 +16,11 @@ solution (t := ‖x‖) solutionEqualsAtom by rfl feasibility (c : by - simp [soCone, Norm.norm]) + unfold soCone + simp [Norm.norm]) optimality by - simp [soCone, Norm.norm] at c ⊢ + unfold soCone at c + simp [Norm.norm] at c ⊢ exact c vconditionElimination diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/PosSemidef.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/PosSemidef.lean index d638e4fb..810fd183 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/PosSemidef.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/PosSemidef.lean @@ -11,9 +11,9 @@ implementationVars implementationObjective Real.Matrix.PSDCone A implementationConstraints solution -solutionEqualsAtom by simp [Real.Matrix.PSDCone] +solutionEqualsAtom by rw [Real.Matrix.PSDCone] feasibility -optimality by simp [Real.Matrix.PSDCone] +optimality by rw [Real.Matrix.PSDCone] vconditionElimination end CvxLean diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean index 95ff4b14..78097d32 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean @@ -10,6 +10,7 @@ namespace CvxLean open Real +set_option trace.Meta.debug true in declare_atom powNegOne [convex] (x : ℝ)- : x ^ (-1) := vconditions (hx : 0 < x) @@ -18,7 +19,7 @@ implementationObjective (t) implementationConstraints (c1 : soCone (t + x) ![t - x, 2]) (c2 : 0 ≤ t) - (c3 : 0 ≤ x) + (c3 : 0 ≤ x) -- TODO: This is not reduced. solution (t := x ^ (-1)) solutionEqualsAtom rfl @@ -55,7 +56,6 @@ vconditionElimination -- NOTE(RFM): It is not straightforward to express it in terms of x ^ (-1) and -- x ^ 2 because of vconditions. -set_option trace.Meta.debug true in declare_atom powNegTwo [convex] (x : ℝ)- : x ^ (-2) := vconditions (hx : 0 < x) @@ -66,7 +66,7 @@ implementationConstraints (c2 : soCone (t₀ + 1) ![t₀ - 1, 2 * t₁]) (c3 : 0 ≤ t₀) (c4 : 0 ≤ t₁) - (c5 : 0 ≤ x) + (c5 : 0 ≤ x) -- TODO: This is not reduced. solution (t₀ := x ^ (-2)) (t₁ := x ^ (-1)) solutionEqualsAtom rfl diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean index 00fcb419..6f1f0e32 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean @@ -17,9 +17,11 @@ solution solutionEqualsAtom rfl feasibility (c1 : by - simp [rotatedSoCone] + unfold rotatedSoCone + simp exact ⟨sq_nonneg x, zero_le_two⟩) optimality by + unfold rotatedSoCone at c1 have h := c1.1 simp at h ⊢ exact h @@ -36,11 +38,12 @@ solution solutionEqualsAtom rfl feasibility (c1 : by - dsimp [Vec.rotatedSoCone] - intros i + unfold Vec.rotatedSoCone + intros t i convert sq.feasibility0 (x i); simp) optimality by intros i + unfold Vec.rotatedSoCone at c1 convert sq.optimality (x i) (t i) (c1 i); simp vconditionElimination diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sqrt.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sqrt.lean index 3e95ef79..8bfa27eb 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sqrt.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sqrt.lean @@ -17,14 +17,20 @@ solutionEqualsAtom by rfl; feasibility (c1 : by - simp [rotatedSoCone] - refine ⟨?_, cond, zero_le_two⟩ + dsimp + unfold rotatedSoCone + refine ⟨?_, cond, by norm_num⟩ + simp rw [sq_sqrt cond]) optimality by intros y hy - simp [rotatedSoCone] at c1 + unfold rotatedSoCone at c1 + simp at c1 have h := c1.1 exact Real.le_sqrt_of_sq_le (le_trans h hy) -vconditionElimination (cond : fun _ hx => c1.2.1.trans hx) +vconditionElimination (cond : by + intros _ hx + unfold rotatedSoCone at c1 + exact c1.2.1.trans hx) end CvxLean diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/XExp.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/XExp.lean index df187aa6..2a6fda06 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/XExp.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/XExp.lean @@ -18,7 +18,7 @@ solution (t₀ := x * exp x) (t₁ := x ^ 2) solutionEqualsAtom rfl feasibility (c1 : by - simp only [expCone] + unfold expCone by_cases (0 < x) . left refine ⟨h, ?_⟩ @@ -30,15 +30,14 @@ feasibility (c2 : by norm_num) (c3 : by simp [posOrthCone, cond]) optimality by { - simp [expCone] at c1 - simp [rotatedSoCone] at c2 + unfold expCone at c1 + simp at c2 cases c1 with | inl c1l => rcases c1l with ⟨hxpos, hxexp⟩ apply le_trans _ hxexp apply mul_le_mul_of_nonneg_left _ c3 - rw [Real.exp_le_exp] - rw [le_div_iff hxpos] + rw [Real.exp_le_exp, le_div_iff hxpos] rw [pow_two] at c2 exact c2 | inr c1r => diff --git a/CvxLean/Tactic/DCP/Atoms.lean b/CvxLean/Tactic/DCP/Atoms.lean index 6cf348d6..72171f0f 100644 --- a/CvxLean/Tactic/DCP/Atoms.lean +++ b/CvxLean/Tactic/DCP/Atoms.lean @@ -86,12 +86,14 @@ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandEla withLocalDeclsD (atomData.impVars.map fun (n, ty) => (n, fun _ => return mkAppNBeta ty xs)) fun vs => do - let originalVarsDecls ← vs.mapM fun v => v.fvarId!.getDecl + let vsDecls ← vs.mapM fun v => v.fvarId!.getDecl + let originalVarsDecls := vsDecls --++ xsDecs let objFun := mkAppNBeta (mkAppNBeta atomData.impObjFun xs) vs let constraints := atomData.impConstrs.map fun c => (`_, mkAppNBeta (mkAppNBeta c xs) vs) return (objFun, constraints, originalVarsDecls) + let xsDecls ← xs.mapM fun x => x.fvarId!.getDecl trace[Meta.debug] "before PAT " let pat ← DCP.mkProcessedAtomTree objCurv objFun constraints.toList originalVarsDecls trace[Meta.debug] "after PAT " From de8c402a924f9b6a68e4609a0da6ddd732b4df59 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 23 Nov 2023 20:08:46 +0000 Subject: [PATCH 005/189] wip: reducing all constraints? --- CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean | 1 + CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean | 18 +++++++++++++++++- .../Tactic/DCP/AtomLibrary/Fns/VecCons.lean | 2 +- CvxLean/Tactic/DCP/Atoms.lean | 17 ++++++++++++----- 4 files changed, 31 insertions(+), 7 deletions(-) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean index ec9f5bce..08471f46 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean @@ -14,6 +14,7 @@ namespace CvxLean open Real +set_option trace.Meta.debug true in declare_atom huber [convex] (x : ℝ)? : huber x := vconditions implementationVars (v : ℝ) (w : ℝ) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean index 6f1f0e32..a042c711 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean @@ -27,12 +27,28 @@ optimality by exact h vconditionElimination +-- TODO: There is an issue matching functions so we need to wrap it. +def test (x : (Fin n) → ℝ) := fun i => ![x i] + +declare_atom vecCons2 [affine] (n : Nat)& (x : (Fin n) → ℝ)? : + test x := +bconditions +homogenity by + unfold test + ext i + simp +additivity by + unfold test + ext i + simp +optimality le_refl _ + declare_atom Vec.sq [convex] (n : ℕ)& (x : Fin n → ℝ)? : x ^ 2 := vconditions implementationVars (t : Fin n → ℝ) implementationObjective (t) implementationConstraints - (c1 : Vec.rotatedSoCone t (fun _ => 1/2) (fun i => ![x i])) + (c1 : Vec.rotatedSoCone t (fun _ => 1/2) (test x)) solution (t := x ^ 2) solutionEqualsAtom rfl diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/VecCons.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/VecCons.lean index 057459f9..f5ee0e98 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/VecCons.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/VecCons.lean @@ -5,7 +5,7 @@ namespace CvxLean open Matrix -declare_atom vec_cons [affine] (n : Nat)& (x : ℝ)+ (y : (Fin n) → ℝ)+ : +declare_atom vecCons [affine] (n : Nat)& (x : ℝ)+ (y : (Fin n) → ℝ)+ : vecCons x y := bconditions homogenity by diff --git a/CvxLean/Tactic/DCP/Atoms.lean b/CvxLean/Tactic/DCP/Atoms.lean index 72171f0f..f27e22dc 100644 --- a/CvxLean/Tactic/DCP/Atoms.lean +++ b/CvxLean/Tactic/DCP/Atoms.lean @@ -87,15 +87,22 @@ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandEla (atomData.impVars.map fun (n, ty) => (n, fun _ => return mkAppNBeta ty xs)) fun vs => do let vsDecls ← vs.mapM fun v => v.fvarId!.getDecl - let originalVarsDecls := vsDecls --++ xsDecs + + let originalVarsDecls := vsDecls let objFun := mkAppNBeta (mkAppNBeta atomData.impObjFun xs) vs let constraints := atomData.impConstrs.map fun c => (`_, mkAppNBeta (mkAppNBeta c xs) vs) return (objFun, constraints, originalVarsDecls) - let xsDecls ← xs.mapM fun x => x.fvarId!.getDecl + let xsDeclsPre ← xs.mapM fun x => x.fvarId!.getDecl + let mut xsDecls := #[] + for i in [:xs.size] do + if atomData.argKinds[i]! != ArgKind.Constant then + xsDecls := xsDecls.push xsDeclsPre[i]! + trace[Meta.debug] "xsDecls: {xsDecls.map (·.userName)}" + trace[Meta.debug] "before PAT " - let pat ← DCP.mkProcessedAtomTree objCurv objFun constraints.toList originalVarsDecls + let pat ← DCP.mkProcessedAtomTree objCurv objFun constraints.toList (xsDecls ++ originalVarsDecls) trace[Meta.debug] "after PAT " -- `pat` is the atom tree resulting from the DCP procedure. @@ -107,7 +114,7 @@ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandEla trace[Meta.debug] "pat.newVarDecls : {pat.newVarDecls.map (·.userName)}" trace[Meta.debug] "pat.newConstrVarsArray : {pat.newConstrVarsArray.map (·.userName)}" - withExistingLocalDecls pat.originalVarsDecls.toList do + withExistingLocalDecls originalVarsDecls.toList do withExistingLocalDecls pat.newVarDecls do withExistingLocalDecls pat.newConstrVarsArray.toList do trace[Meta.debug] "pat opt: {pat.optimality}" @@ -115,7 +122,7 @@ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandEla trace[Meta.debug] "pat opt constr: {c}" check c -- `vs1` are the variables already present in the unreduced graph implementation. - let vs1 := pat.originalVarsDecls.map (mkFVar ·.fvarId) + let vs1 := originalVarsDecls.map (mkFVar ·.fvarId) -- `vs2` are the variables to be added to the graph implementation. let vs2 := pat.newVarDecls.toArray.map (mkFVar ·.fvarId) let vs1Sol := atomData.solution.map fun s => mkAppNBeta s xs From e851e6c6bcccf4ae41aaec20c1faf52eb379dea8 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 23 Nov 2023 20:13:31 +0000 Subject: [PATCH 006/189] fix: extra constraints in Huber --- CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean index 08471f46..2ef9ff43 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean @@ -109,12 +109,15 @@ optimality by intros i simp [Vec.huber] rw [←rpow_two] - apply huber.optimality (x i) (v i) (w i) ((w i) ^ 2) + apply huber.optimality (x i) (v i) (w i) ((w i) ^ 2) (|x i|) { unfold posOrthCone; simpa using c1 i } { unfold posOrthCone; simpa using c2 i } { unfold posOrthCone; simpa using c3 i } { unfold posOrthCone; simpa using c4 i } { unfold rotatedSoCone; simp [sq_nonneg]; norm_num } + { unfold posOrthCone; simp [le_abs_self] } + { unfold posOrthCone; rw [←sub_le_iff_le_add, zero_sub]; simp [neg_le_abs_self] } + vconditionElimination end CvxLean From 37a1cd291ad4cc4deea247ea883c635abd512787 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Sat, 25 Nov 2023 17:39:45 +0000 Subject: [PATCH 007/189] wip: optimality for eq needs an intros --- CvxLean/Tactic/DCP/AtomLibrary/Sets/Eq.lean | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Sets/Eq.lean b/CvxLean/Tactic/DCP/AtomLibrary/Sets/Eq.lean index 737d605f..53d5439d 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Sets/Eq.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Sets/Eq.lean @@ -16,9 +16,10 @@ solutionEqualsAtom by exact Iff.intro Eq.symm Eq.symm feasibility optimality by - unfold Real.zeroCone - simp [sub_eq_iff_eq_add, zero_add] - exact Eq.symm + intros h + unfold Real.zeroCone at h + simp [sub_eq_iff_eq_add, zero_add] at h + exact h.symm vconditionElimination end CvxLean From 79f1a9bf7780847bb68623743d6b213c06d06fe7 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Sat, 25 Nov 2023 17:40:33 +0000 Subject: [PATCH 008/189] wip: adjust optimality condition for non-trivial atoms --- CvxLean/Tactic/DCP/Atoms.lean | 134 +++++++++++++++++++++++++--------- 1 file changed, 99 insertions(+), 35 deletions(-) diff --git a/CvxLean/Tactic/DCP/Atoms.lean b/CvxLean/Tactic/DCP/Atoms.lean index f27e22dc..152f2fe5 100644 --- a/CvxLean/Tactic/DCP/Atoms.lean +++ b/CvxLean/Tactic/DCP/Atoms.lean @@ -74,6 +74,44 @@ where } return mkConst name + +/-- -/ +def withCopyOfMonoXs [Inhabited α] (xs : Array Expr) (argKinds : Array ArgKind) (f : Array Expr → Array Expr → Array ArgKind → TermElabM α) : TermElabM α := do + -- Determine subset of monotone arguments and their behavior. + let mut argDeclInfo : Array (Lean.Name × (Array Expr → TermElabM Expr)) := #[] + let mut monoXs := #[] + let mut monoArgKind := #[] + for i in [:xs.size] do + if argKinds[i]! != ArgKind.Constant ∧ argKinds[i]! != ArgKind.Neither then + let ty := ← inferType xs[i]! + let name := (← FVarId.getDecl xs[i]!.fvarId!).userName + argDeclInfo := argDeclInfo.push (name, fun _ => pure ty) + monoXs := monoXs.push xs[i]! + monoArgKind := monoArgKind.push argKinds[i]! + + withLocalDeclsD argDeclInfo fun ys => do + return ← f monoXs ys monoArgKind + +/-- -/ +def shiftingArgs (curv : Curvature) (xs : Array Expr) (argKinds : Array ArgKind) (mkConcl : Array Expr → Array Expr → TermElabM Expr) : TermElabM Expr := + withCopyOfMonoXs xs argKinds fun monoXs ys monoArgKind => do + let mut ty := ← mkConcl monoXs ys + for i' in [:monoXs.size] do + let i := monoXs.size - 1 - i' + ty ← match curvatureInArg curv monoArgKind[i]! with + | Curvature.Concave => mkArrow (← mkLe monoXs[i]! ys[i]!) ty + | Curvature.Convex => mkArrow (← mkLe ys[i]! monoXs[i]!) ty + | _ => throwError "invalid curvature" + return ← mkForallFVars ys ty + +/-- TODO: just the number of relevant opt args. -/ +def getMonoArgsCount (curv : Curvature) (argKinds : Array ArgKind) : ℕ := + argKinds.foldl (fun acc argKind => + match curvatureInArg curv argKind with + | Curvature.Concave => 1 + acc + | Curvature.Convex => 1 + acc + | _ => acc) 0 + /-- Use the DCP procedure to reduce the graph implementation to a problem that uses only cone constraints and affine atoms. -/ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandElabM GraphAtomData := do @@ -126,6 +164,10 @@ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandEla -- `vs2` are the variables to be added to the graph implementation. let vs2 := pat.newVarDecls.toArray.map (mkFVar ·.fvarId) let vs1Sol := atomData.solution.map fun s => mkAppNBeta s xs + trace[Meta.debug] "vs1: {vs1}" + trace[Meta.debug] "vs2: {vs2}" + trace[Meta.debug] "vs1Sol: {vs1Sol}" + trace[Meta.debug] "xs: {xs}" -- TODO: move because copied from DCP tactic. let reducedConstrs := pat.reducedExprs.constr.map Tree.val @@ -141,7 +183,6 @@ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandEla for ncp in newConstrProofs do trace[Meta.debug] "newConstrProofs: {← inferType ncp}" - -- let test := ← atomData.feasibility.mapM (fun e => -- Meta.forallTelescope e (fun xs a => do -- trace[Meta.debug] "a: {←inferType a}")) @@ -194,16 +235,65 @@ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandEla throwError ("Expected same length: {reducedConstrs} and " ++ "{constraintsFromReducedConstraints}") - let newOptimality := mkAppN atomData.optimality (xs ++ vs1) + let objFunFromReducedObjFun := pat.optimality.objFun.val + trace[Meta.debug] "objFunFromReducedObjFun: {← inferType objFunFromReducedObjFun}" + + trace[Meta.debug] "pat.optimality.objFun: {← inferType atomData.optimality}" + let optimalityXs := mkAppN atomData.optimality (xs ++ vs1) + trace[Meta.debug] "newOptimality: {← inferType optimalityXs}" let newOptimality ← withLocalDeclsDNondep (reducedConstrs.map (fun rc => (`_, rc))) fun cs => do - let mut newOptimality := newOptimality + -- First, apply all constraints. + let mut optimalityAfterReduced := optimalityXs for i in [:reducedConstrs.size] do let c := mkApp constraintsFromReducedConstraints[i]! cs[i]! - newOptimality := mkApp newOptimality c + optimalityAfterReduced := mkApp optimalityAfterReduced c + -- Then, adjust the condition using `objFunFromReducedObjFun`. + let (condTy, condY, condZ) ← ((do + match (← inferType objFunFromReducedObjFun).le? with + | some (t, y, z) => return (t, y, z) + | none => throwError "Cannot build optimality proof") : MetaM _) + + let monoArgsCount := getMonoArgsCount objCurv atomData.argKinds + let optimalityAfterApplyWithConditionAdjusted ← + lambdaTelescope (← whnf optimalityAfterReduced) <| fun xs e => do + -- Every extra argument has an extra condition, e.g. x', x ≤ x. + let monoArgs := xs[:2 * monoArgsCount] + let optCondition ← mkLambdaFVars xs[2 * monoArgsCount:] e + let newCond ← + if atomData.curvature == Curvature.Convex then + mkAppOptM ``le_trans #[ + condTy, none, none, condY, condZ, + optCondition, objFunFromReducedObjFun] + else + -- TODO: concave. but convex_set too? + mkAppOptM ``le_trans #[ + condTy, none, condY, condZ, none, + objFunFromReducedObjFun, optCondition] + mkLambdaFVars monoArgs newCond + + trace[Meta.debug] "optimalityAfterApplyWithConditionAdjusted: {← inferType optimalityAfterApplyWithConditionAdjusted}" + trace[Meta.debug] "newOptimality applied: {← inferType optimalityAfterReduced}" let ds := pat.newConstrVarsArray.map (mkFVar ·.fvarId) - mkLambdaFVars (xs ++ vs1 ++ vs2) <| ← mkLambdaFVars (cs ++ ds) newOptimality - + mkLambdaFVars (xs ++ vs1 ++ vs2) <| ← mkLambdaFVars (cs ++ ds) optimalityAfterApplyWithConditionAdjusted + -- trace[Meta.debug] "newOptimality2: {← inferType newOptimality2}" + + -- -- NOTE: new optimality might be of the form ∀ i. x i ≤ y i, so + -- -- we cannot apply le_trans directly. + -- let x ← lambdaTelescope (← whnf newOptimality2) (fun xs a => do + -- trace[Meta.debug] "a: {← inferType a}" + -- trace[Meta.debug] "xs: {xs}" + -- return ()) + -- match (← inferType objFunFromReducedObjFun).le? with + -- | some (t, y, z) => + -- trace[Meta.debug] "atomData.argKinds: {atomData.argKinds}" + -- let x ← getOptArgs objCurv xs atomData.argKinds + -- trace[Meta.debug] "x: {x}" + -- -- mkAppOptM ``le_trans #[ + -- -- t, none, none, y, z, + -- -- newOptimality2, objFunFromReducedObjFun] + -- let ds := pat.newConstrVarsArray.map (mkFVar ·.fvarId) + -- mkLambdaFVars (xs ++ vs1 ++ vs2) <| ← mkLambdaFVars (cs ++ ds) newOptimality2 --optimalityAdjusted let mut newVCondElims := #[] for vcondElim in atomData.vcondElim do let newVCondElim := mkAppN vcondElim (xs ++ vs1) @@ -222,7 +312,7 @@ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandEla (← pat.newVarDecls.toArray.mapM fun decl => do return (decl.userName, ← mkLambdaFVars xs decl.type)) impObjFun := ← mkLambdaFVars xs $ ← mkLambdaFVars (vs1 ++ vs2) $ - pat.reducedExprs.objFun.val + pat.reducedExprs.objFun.val impConstrs := ← (reducedConstrs ++ pat.newConstrs).mapM (fun c => do return ← mkLambdaFVars xs $ ← mkLambdaFVars (vs1 ++ vs2) $ c) solution := atomData.solution.append @@ -412,34 +502,7 @@ def withCopyOfNonConstVars (xs : Array Expr) (argKinds : Array ArgKind) (f : Arr j := j + 1 return ← f ys allYs -/-- -/ -def withCopyOfMonoXs (xs : Array Expr) (argKinds : Array ArgKind) (f : Array Expr → Array Expr → Array ArgKind → TermElabM Expr) : TermElabM Expr := do - -- Determine subset of monotone arguments and their behavior. - let mut argDeclInfo : Array (Lean.Name × (Array Expr → TermElabM Expr)) := #[] - let mut monoXs := #[] - let mut monoArgKind := #[] - for i in [:xs.size] do - if argKinds[i]! != ArgKind.Constant ∧ argKinds[i]! != ArgKind.Neither then - let ty := ← inferType xs[i]! - let name := (← FVarId.getDecl xs[i]!.fvarId!).userName - argDeclInfo := argDeclInfo.push (name, fun _ => pure ty) - monoXs := monoXs.push xs[i]! - monoArgKind := monoArgKind.push argKinds[i]! - withLocalDeclsD argDeclInfo fun ys => do - return ← f monoXs ys monoArgKind - -/-- -/ -def shiftingArgs (curv : Curvature) (xs : Array Expr) (argKinds : Array ArgKind) (mkConcl : Array Expr → Array Expr → TermElabM Expr) : TermElabM Expr := - withCopyOfMonoXs xs argKinds fun monoXs ys monoArgKind => do - let mut ty := ← mkConcl monoXs ys - for i' in [:monoXs.size] do - let i := monoXs.size - 1 - i' - ty ← match curvatureInArg curv monoArgKind[i]! with - | Curvature.Concave => mkArrow (← mkLe monoXs[i]! ys[i]!) ty - | Curvature.Convex => mkArrow (← mkLe ys[i]! monoXs[i]!) ty - | _ => throwError "invalid curvature" - return ← mkForallFVars ys ty /-- -/ def elabOpt (curv : Curvature) (argDecls : Array LocalDecl) (expr : Expr) (bconds : Array (Lean.Name × Expr)) @@ -460,8 +523,9 @@ def elabOpt (curv : Curvature) (argDecls : Array LocalDecl) (expr : Expr) (bcond | Curvature.Concave => return ← mkLe impObj body | Curvature.Convex => return ← mkLe body impObj | _ => throwError "invalid curvature: {curv}" - trace[Meta.debug] "{ty}" + trace[Meta.debug] "elabOpt ensuring {ty}" let opt ← Elab.Term.elabTermAndSynthesizeEnsuringType stx (some ty) + trace[Meta.debug] "elabOpt opt: {← inferType opt}" return ← mkLambdaFVars xs $ ← mkLambdaFVars as $ ← mkLambdaFVars vs $ ← mkLambdaFVars is opt /-- -/ From 87899dae978f9fcca7743932e5f9fe60e43f6d81 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Sat, 25 Nov 2023 18:18:52 +0000 Subject: [PATCH 009/189] fix: use atom range to build inequality --- CvxLean/Tactic/DCP/Atoms.lean | 13 +++++-------- 1 file changed, 5 insertions(+), 8 deletions(-) diff --git a/CvxLean/Tactic/DCP/Atoms.lean b/CvxLean/Tactic/DCP/Atoms.lean index 152f2fe5..61cddf42 100644 --- a/CvxLean/Tactic/DCP/Atoms.lean +++ b/CvxLean/Tactic/DCP/Atoms.lean @@ -117,7 +117,8 @@ uses only cone constraints and affine atoms. -/ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandElabM GraphAtomData := do liftTermElabM do -- `xs` are the arguments of the atom. - lambdaTelescope atomData.expr fun xs _ => do + lambdaTelescope atomData.expr fun xs atomExpr => do + let atomRange ← inferType atomExpr -- Call DCP on graph implementation. let (objFun, constraints, originalVarsDecls) ← @@ -249,11 +250,7 @@ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandEla let c := mkApp constraintsFromReducedConstraints[i]! cs[i]! optimalityAfterReduced := mkApp optimalityAfterReduced c -- Then, adjust the condition using `objFunFromReducedObjFun`. - let (condTy, condY, condZ) ← ((do - match (← inferType objFunFromReducedObjFun).le? with - | some (t, y, z) => return (t, y, z) - | none => throwError "Cannot build optimality proof") : MetaM _) - + trace[Meta.debug] "optimalityAfterReduced: {← inferType optimalityAfterReduced}" let monoArgsCount := getMonoArgsCount objCurv atomData.argKinds let optimalityAfterApplyWithConditionAdjusted ← lambdaTelescope (← whnf optimalityAfterReduced) <| fun xs e => do @@ -263,12 +260,12 @@ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandEla let newCond ← if atomData.curvature == Curvature.Convex then mkAppOptM ``le_trans #[ - condTy, none, none, condY, condZ, + atomRange, none, none, none, none, optCondition, objFunFromReducedObjFun] else -- TODO: concave. but convex_set too? mkAppOptM ``le_trans #[ - condTy, none, condY, condZ, none, + atomRange, none, none, none, none, objFunFromReducedObjFun, optCondition] mkLambdaFVars monoArgs newCond From b4b1913292c7dcf33cc300beb9ec33316dfa4b01 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Sat, 25 Nov 2023 18:29:52 +0000 Subject: [PATCH 010/189] wip: issue with `replaceProjections` --- CvxLean/Meta/Minimization.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/CvxLean/Meta/Minimization.lean b/CvxLean/Meta/Minimization.lean index 2913f3ce..a84da323 100644 --- a/CvxLean/Meta/Minimization.lean +++ b/CvxLean/Meta/Minimization.lean @@ -108,6 +108,7 @@ where MetaM (Option Expr) := do /- `first` tells us whether the outermost projection was `.1` (`some true`) or `.2` (`some false`). If this is not a recursive call, `first` is `none`. -/ + -- TODO: This won't recognize "p.2 i". match first, e, rs with | _, Expr.fvar fVarId, [r] => if fVarId == p then return r else return none From 69764b807d23f1f688f7d77d8287be6acd926549 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Sat, 25 Nov 2023 18:33:26 +0000 Subject: [PATCH 011/189] test: recovered most of the old DCP tests --- CvxLean/Test/DCP/Dcp.lean | 134 ++++++++++++++++++++------------------ 1 file changed, 69 insertions(+), 65 deletions(-) diff --git a/CvxLean/Test/DCP/Dcp.lean b/CvxLean/Test/DCP/Dcp.lean index 7764acfd..e1599c9a 100644 --- a/CvxLean/Test/DCP/Dcp.lean +++ b/CvxLean/Test/DCP/Dcp.lean @@ -58,61 +58,66 @@ noncomputable def test001 : Solution $ dcp sorry --- noncomputable def test002 : Solution $ --- optimization (x y : ℝ) --- minimize exp (huber y) --- subject to --- c0 : exp (exp (huber x)) ≤ y --- := by --- dcp --- sorry - --- noncomputable def test003 : Solution $ --- optimization (x y : ℝ) --- minimize (2 : ℝ) * (huber (y + x)) --- subject to --- c0 : x ≤ y --- := by --- dcp --- sorry - --- noncomputable def testVec0 [Fintype m] : Solution $ --- optimization (x y : m → ℝ) --- minimize (0 : ℝ) --- subject to --- c0 : Vec.exp y ≤ x --- := by --- dcp --- sorry - --- noncomputable def testVec [Fintype m] : Solution $ --- optimization (x y : m → ℝ) --- minimize (0 : ℝ) --- subject to --- c0 : Vec.exp (Vec.exp x) ≤ x --- := by --- dcp --- sorry - --- noncomputable def test_Vec_huber {n : ℕ} : Solution $ --- optimization (x y : Fin n → ℝ) --- minimize (0 : ℝ) --- subject to --- c0 : Vec.huber x ≤ x --- := by --- dcp --- sorry - --- noncomputable def test_Vec_kl_div {n : ℕ} : Solution $ --- optimization (x y : Fin n → ℝ) --- minimize (0 : ℝ) --- subject to --- cx : 0 ≤ x --- cy : ∀ i, 0 < y i --- c0 : Vec.klDiv x y ≤ x --- := by --- dcp --- sorry +noncomputable def test002 : Solution $ + optimization (x y : ℝ) + minimize exp (huber y) + subject to + c0 : exp (exp (huber x)) ≤ y +:= by + dcp + sorry + +noncomputable def test003 : Solution $ + optimization (x y : ℝ) + minimize (2 : ℝ) * (huber (y + x)) + subject to + c0 : x ≤ y +:= by + dcp + sorry + +-- TODO: any fin type +noncomputable def testVec0 : Solution $ + optimization (x y : Fin m → ℝ) + minimize (0 : ℝ) + subject to + c0 : Vec.exp y ≤ x +:= by + dcp + sorry + +-- TODO: any fin type +noncomputable def testVec : Solution $ + optimization (x y : Fin m → ℝ) + minimize (0 : ℝ) + subject to + c0 : Vec.exp (Vec.exp x) ≤ x +:= by + dcp + sorry + +noncomputable def test_Vec_huber {n : ℕ} : Solution $ +optimization (x y : Fin n → ℝ) + minimize (0 : ℝ) + subject to + c0 : Vec.huber x ≤ x +:= by + dcp + sorry + +-- TODO: notation. +-- TODO: error if the constraint is ∀ i +noncomputable def test_Vec_kl_div {n : ℕ} : Solution $ +optimization (x y : Fin n → ℝ) + minimize (0 : ℝ) + subject to + cx : 0 ≤ x + cy : StrongLT 0 y + -- cy : ∀ i, 0 < (y) i + c0 : Vec.klDiv x y ≤ x +:= by + dcp + sorry noncomputable def test2 : Solution $ optimization (x y : ℝ) @@ -152,8 +157,6 @@ noncomputable example sorry - - noncomputable def a1 : ℝ := 3 noncomputable def b1 : ℝ := 4 noncomputable def c1 : ℝ := 5 @@ -162,11 +165,12 @@ noncomputable def d1 : ℝ := 6 set_option trace.Elab.debug true -- TODO: --- noncomputable example : Solution $ --- minimization! (x y : ℝ) : --- objective (x) --- constraints --- (hlog : 0 < exp a1) --- (e : exp y ≤ log (exp a1)) --- (hsqrt : 0 ≤ x) := by --- dcp +noncomputable example : Solution $ + optimization (x y : ℝ) + minimize (x) + subject to + hlog : 0 ≤ exp a1 + e : exp y ≤ log (exp a1) + hsqrt : 0 ≤ x := by + dcp -- TODO: first constraint not reduced. We need to handle constant constraints. + sorry From c4b63550012cb31f40d47abd9670d3cf14b60771 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Sat, 25 Nov 2023 18:39:11 +0000 Subject: [PATCH 012/189] fix: make atoms reducible for adjusted optimality --- CvxLean/Lib/Cones/ExpCone.lean | 3 --- CvxLean/Lib/Cones/PSDCone.lean | 1 - CvxLean/Lib/Cones/PosOrthCone.lean | 3 --- CvxLean/Lib/Cones/SOCone.lean | 4 ---- CvxLean/Lib/Cones/ZeroCone.lean | 3 --- CvxLean/Tactic/DCP/AtomLibrary/Sets/Eq.lean | 7 +++---- 6 files changed, 3 insertions(+), 18 deletions(-) diff --git a/CvxLean/Lib/Cones/ExpCone.lean b/CvxLean/Lib/Cones/ExpCone.lean index fc7e3551..7ae33c2d 100644 --- a/CvxLean/Lib/Cones/ExpCone.lean +++ b/CvxLean/Lib/Cones/ExpCone.lean @@ -2,15 +2,12 @@ import Mathlib.Data.Complex.Exponential namespace Real -@[irreducible] def expCone (x y z : Real) : Prop := (0 < y ∧ y * exp (x / y) ≤ z) ∨ (y = 0 ∧ 0 ≤ z ∧ x ≤ 0) -@[irreducible] def Vec.expCone (x y z : Fin n → Real) : Prop := ∀ i, Real.expCone (x i) (y i) (z i) -@[irreducible] theorem exp_iff_expCone (t x : Real) : exp x ≤ t ↔ expCone x 1 t := by unfold expCone rw [iff_def] diff --git a/CvxLean/Lib/Cones/PSDCone.lean b/CvxLean/Lib/Cones/PSDCone.lean index 294a78d4..aed62387 100644 --- a/CvxLean/Lib/Cones/PSDCone.lean +++ b/CvxLean/Lib/Cones/PSDCone.lean @@ -3,7 +3,6 @@ import Mathlib.LinearAlgebra.Matrix.PosDef namespace Real -@[irreducible] def Matrix.PSDCone {m} [Fintype m] (M : Matrix m m Real) : Prop := Matrix.PosSemidef M diff --git a/CvxLean/Lib/Cones/PosOrthCone.lean b/CvxLean/Lib/Cones/PosOrthCone.lean index 730ff49b..4e28667e 100644 --- a/CvxLean/Lib/Cones/PosOrthCone.lean +++ b/CvxLean/Lib/Cones/PosOrthCone.lean @@ -3,15 +3,12 @@ import CvxLean.Lib.Math.Data.Matrix namespace Real -@[irreducible] def posOrthCone (x : Real) : Prop := 0 ≤ x -@[irreducible] def Vec.posOrthCone (x : Fin n → Real) : Prop := ∀ i, Real.posOrthCone (x i) -@[irreducible] def Matrix.posOrthCone (M : Matrix (Fin m) (Fin n) Real) : Prop := ∀ i j, Real.posOrthCone (M i j) diff --git a/CvxLean/Lib/Cones/SOCone.lean b/CvxLean/Lib/Cones/SOCone.lean index ad6ace6c..81c22717 100644 --- a/CvxLean/Lib/Cones/SOCone.lean +++ b/CvxLean/Lib/Cones/SOCone.lean @@ -14,19 +14,15 @@ open BigOperators variable [Fintype m] [Fintype n] -@[irreducible] def soCone (t : Real) (x : n → Real) : Prop := sqrt (∑ i, x i ^ 2) ≤ t -@[irreducible] def rotatedSoCone (v w : Real) (x : n → Real) : Prop := (∑ i, x i ^ 2) ≤ (v * w) * 2 ∧ 0 ≤ v ∧ 0 ≤ w -@[irreducible] def Vec.soCone (t : m → Real) (X : Matrix m n Real) : Prop := ∀ i, Real.soCone (t i) (X i) -@[irreducible] def Vec.rotatedSoCone (v w : m → Real) (X : Matrix m n Real) : Prop := ∀ i, Real.rotatedSoCone (v i) (w i) (X i) diff --git a/CvxLean/Lib/Cones/ZeroCone.lean b/CvxLean/Lib/Cones/ZeroCone.lean index 1fc22dca..15d27ce5 100644 --- a/CvxLean/Lib/Cones/ZeroCone.lean +++ b/CvxLean/Lib/Cones/ZeroCone.lean @@ -3,15 +3,12 @@ import CvxLean.Lib.Math.Data.Matrix namespace Real -@[irreducible] def zeroCone (x : Real) : Prop := x = 0 -@[irreducible] def Vec.zeroCone (x : Fin n → Real) : Prop := ∀ i, Real.zeroCone (x i) -@[irreducible] def Matrix.zeroCone (M : Matrix (Fin n) (Fin m) Real) : Prop := ∀ i j, Real.zeroCone (M i j) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Sets/Eq.lean b/CvxLean/Tactic/DCP/AtomLibrary/Sets/Eq.lean index 53d5439d..737d605f 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Sets/Eq.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Sets/Eq.lean @@ -16,10 +16,9 @@ solutionEqualsAtom by exact Iff.intro Eq.symm Eq.symm feasibility optimality by - intros h - unfold Real.zeroCone at h - simp [sub_eq_iff_eq_add, zero_add] at h - exact h.symm + unfold Real.zeroCone + simp [sub_eq_iff_eq_add, zero_add] + exact Eq.symm vconditionElimination end CvxLean From e94ae636ff66c2df34bd41c09f8498ab33123f69 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 4 Dec 2023 18:04:00 -0500 Subject: [PATCH 013/189] test: DCP exercises from course --- rewriter/src/tests.rs | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/rewriter/src/tests.rs b/rewriter/src/tests.rs index a03f5e33..4e493de0 100644 --- a/rewriter/src/tests.rs +++ b/rewriter/src/tests.rs @@ -239,4 +239,24 @@ fn test_log_div_rev_obj() { ]); } +// More + +#[test] +fn test_3_32() { + assert_steps_with_domain( + vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], + "(neg (div 1 (mul (var x) (var y))))", + vec![ + ]); +} + +#[test] +fn test_3_33() { + assert_steps_with_domain( + vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], + "(sqrt (add 1 (div (pow (var x) 4) (var y))))", + vec![ + ]); +} + } From e3a5a282790127ea26ea1b00ff9a32e12e5fcb79 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 4 Dec 2023 18:50:41 -0500 Subject: [PATCH 014/189] wip: `geo_mean` and `quad_over_lin` in cost check --- rewriter/src/cost.rs | 114 +++++++++++++++++++++++++++++++++---------- 1 file changed, 88 insertions(+), 26 deletions(-) diff --git a/rewriter/src/cost.rs b/rewriter/src/cost.rs index 02bfc8ce..7d735d05 100644 --- a/rewriter/src/cost.rs +++ b/rewriter/src/cost.rs @@ -1,3 +1,5 @@ +use std::cmp::Ordering; + use egg::{*}; use crate::domain; @@ -13,10 +15,19 @@ pub struct DCPCost<'a> { pub egraph: &'a optimization::EGraph, } +#[derive(Debug, Clone, PartialEq)] +pub struct OptimizationWrapper(Optimization); + +impl PartialOrd for OptimizationWrapper { + fn partial_cmp(&self, other: &Self) -> Option { + Some(Ordering::Equal) + } +} + impl<'a> CostFunction for DCPCost<'a> { // Curvature + number of variables (with repetition) + term size. // In lexicographic order. - type Cost = (Curvature, u32, u32); + type Cost = (Curvature, u32, u32, OptimizationWrapper); fn cost(&mut self, enode: &Optimization, mut costs: C) -> Self::Cost where C: FnMut(Id) -> Self::Cost @@ -30,6 +41,9 @@ impl<'a> CostFunction for DCPCost<'a> { macro_rules! get_term_size { ($i:expr) => { costs(*$i).2 } } + macro_rules! get_node { + ($i:expr) => { costs(*$i).3 } + } let get_domain = |i: &Id| self.egraph[*i].data.domain.clone(); @@ -111,7 +125,25 @@ impl<'a> CostFunction for DCPCost<'a> { term_size = 1 + get_term_size!(a); } Optimization::Sqrt(a) => { - curvature = curvature::of_concave_increasing_fn(get_curvature!(a)); + // TODO: Temporary. geo_mean. + let mut is_geo_mean = false; + match get_node!(a).0 { + Optimization::Mul([b, c]) => { + let affine = + get_curvature!(&b) == Curvature::Affine && + get_curvature!(&c) == Curvature::Affine; + let pos = + domain::option_is_pos(get_domain(&b).as_ref()) && + domain::option_is_pos(get_domain(&c).as_ref()); + is_geo_mean = affine && pos; + } + _ => {} + } + if is_geo_mean { + curvature = Curvature::Convex; + } else { + curvature = curvature::of_concave_increasing_fn(get_curvature!(a)); + } num_vars = get_num_vars!(a); term_size = 1 + get_term_size!(a); } @@ -154,35 +186,64 @@ impl<'a> CostFunction for DCPCost<'a> { term_size = 1 + get_term_size!(a) + get_term_size!(b); } Optimization::Div([a, b]) => { - let db_o = get_domain(b); - curvature = match (get_is_constant(a), get_is_constant(b)) { - (true, true) => { - match db_o { - Some(db) => { - if domain::does_not_contain_zero(&db) { - Curvature::Constant - } else { - Curvature::Unknown + // TODO: Temporary. quad_over_lin. + let mut is_quad_over_lin = false; + match get_node!(a).0 { + Optimization::Pow([c, d]) => { + let curvature_check = + get_curvature!(b) == Curvature::Affine && + get_curvature!(&c) == Curvature::Affine && + get_curvature!(&d) == Curvature::Constant; + let pos_check = + domain::option_is_pos(get_domain(&b).as_ref()); + let pow_two_check = + match get_domain(&d) { + Some(dd) => + match domain::Domain::get_constant(&dd) { + Some (dd_f) => dd_f == 2.0, + _ => false + } + _ => false + }; + is_quad_over_lin = curvature_check && pos_check && pow_two_check; + } + _ => {} + } + + if is_quad_over_lin { + println!("quad_over_lin"); + curvature = Curvature::Convex; + } else { + let db_o = get_domain(b); + curvature = match (get_is_constant(a), get_is_constant(b)) { + (true, true) => { + match db_o { + Some(db) => { + if domain::does_not_contain_zero(&db) { + Curvature::Constant + } else { + Curvature::Unknown + } + } - + None => { Curvature::Unknown } } - None => { Curvature::Unknown } } - } - (false, true) => { - match db_o { - Some(db) => { - if domain::does_not_contain_zero(&db) { - curvature::of_mul_by_const(get_curvature!(a), db) - } else { - Curvature::Unknown + (false, true) => { + match db_o { + Some(db) => { + if domain::does_not_contain_zero(&db) { + curvature::of_mul_by_const(get_curvature!(a), db) + } else { + Curvature::Unknown + } } + None => { Curvature::Unknown } } - None => { Curvature::Unknown } } - } - _ => { Curvature::Unknown } - }; + _ => { Curvature::Unknown } + }; + } num_vars = get_num_vars!(a) + get_num_vars!(b); term_size = 1 + get_term_size!(a) + get_term_size!(b); } @@ -230,6 +291,7 @@ impl<'a> CostFunction for DCPCost<'a> { } } - return (curvature, num_vars, term_size); + return (curvature, num_vars, term_size, OptimizationWrapper(enode.clone())); } + } From f54b8641d54aa624893afd64c9e3d7470e1bf271 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 4 Dec 2023 18:51:00 -0500 Subject: [PATCH 015/189] test: `geo_mean` and `quad_over_lin` --- rewriter/src/tests.rs | 20 +++++++++++++++++++- 1 file changed, 19 insertions(+), 1 deletion(-) diff --git a/rewriter/src/tests.rs b/rewriter/src/tests.rs index 4e493de0..31726fc3 100644 --- a/rewriter/src/tests.rs +++ b/rewriter/src/tests.rs @@ -241,11 +241,29 @@ fn test_log_div_rev_obj() { // More +#[test] +fn test_geo_mean() { + assert_steps_with_domain( + vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], + "(sqrt (mul (var x) (var y)))", + vec![ + ]); +} + +#[test] +fn test_quad_over_lin() { + assert_steps_with_domain( + vec![("x", domain::free_dom()), ("y", domain::pos_dom())], + "(div (pow (var x) 2) (var y))", + vec![ + ]); +} + #[test] fn test_3_32() { assert_steps_with_domain( vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], - "(neg (div 1 (mul (var x) (var y))))", + "(div 1 (mul (var x) (var y)))", vec![ ]); } From f8b0ab5296c732aca4335659f16298af9da78451 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 4 Dec 2023 19:05:02 -0500 Subject: [PATCH 016/189] wip: add `norm2` to handle example 3.33 --- rewriter/src/cost.rs | 40 +++++++++++++++++++++++++++++++++++++--- rewriter/src/tests.rs | 17 +++++++++++++++-- 2 files changed, 52 insertions(+), 5 deletions(-) diff --git a/rewriter/src/cost.rs b/rewriter/src/cost.rs index 7d735d05..91322a31 100644 --- a/rewriter/src/cost.rs +++ b/rewriter/src/cost.rs @@ -139,7 +139,41 @@ impl<'a> CostFunction for DCPCost<'a> { } _ => {} } - if is_geo_mean { + // TODO: Temporary. norm2. + let mut is_norm2 = false; + match get_node!(a).0 { + Optimization::Add([b, c]) => { + match (get_node!(&b).0, get_node!(&c).0) { + (Optimization::Pow([d, e]), Optimization::Pow([f, g])) => { + let curvature_check = + get_curvature!(&d) <= Curvature::Convex && + get_curvature!(&f) <= Curvature::Convex; + let pow_two_first = + match get_domain(&e) { + Some(de) => + match domain::Domain::get_constant(&de) { + Some (de_f) => de_f == 2.0, + _ => false + } + _ => false + }; + let pow_two_second = + match get_domain(&g) { + Some(dg) => + match domain::Domain::get_constant(&dg) { + Some (dg_f) => dg_f == 2.0, + _ => false + } + _ => false + }; + is_norm2 = curvature_check && pow_two_first && pow_two_second; + } + _ => {} + } + } + _ => {} + } + if is_geo_mean || is_norm2 { curvature = Curvature::Convex; } else { curvature = curvature::of_concave_increasing_fn(get_curvature!(a)); @@ -191,8 +225,8 @@ impl<'a> CostFunction for DCPCost<'a> { match get_node!(a).0 { Optimization::Pow([c, d]) => { let curvature_check = - get_curvature!(b) == Curvature::Affine && - get_curvature!(&c) == Curvature::Affine && + get_curvature!(b) <= Curvature::Concave && + get_curvature!(&c) <= Curvature::Affine && get_curvature!(&d) == Curvature::Constant; let pos_check = domain::option_is_pos(get_domain(&b).as_ref()); diff --git a/rewriter/src/tests.rs b/rewriter/src/tests.rs index 31726fc3..ba5f11e7 100644 --- a/rewriter/src/tests.rs +++ b/rewriter/src/tests.rs @@ -259,6 +259,15 @@ fn test_quad_over_lin() { ]); } +#[test] +fn test_norm2() { + assert_steps_with_domain( + vec![("x", domain::free_dom()), ("y", domain::free_dom())], + "(sqrt (add (pow (var x) 2) (pow (var y) 2)))", + vec![ + ]); +} + #[test] fn test_3_32() { assert_steps_with_domain( @@ -270,9 +279,13 @@ fn test_3_32() { #[test] fn test_3_33() { + // NOTE(RFM): The interesting thing here is that we need to somehow have + // 4 -> 2 * 2. Moreover, getting the simplest nodes when matching for + // norm is not enough! assert_steps_with_domain( - vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], - "(sqrt (add 1 (div (pow (var x) 4) (var y))))", + vec![("x", domain::free_dom()), ("y", domain::pos_dom())], + // "(sqrt (add 1 (div (pow (var x) 4) (var y))))", + "(sqrt (add 1 (div (pow (var x) (mul 2 2)) (var y))))", vec![ ]); } From f71b4af662561ea61ddbdd83c5f23c83cb0e43e3 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 09:08:03 -0500 Subject: [PATCH 017/189] feat: extend egg language with `xexp`, `entr`, `qol` and `norm2` --- rewriter/src/optimization.rs | 8 ++++++-- rewriter/src/rules.rs | 11 +++++++++++ 2 files changed, 17 insertions(+), 2 deletions(-) diff --git a/rewriter/src/optimization.rs b/rewriter/src/optimization.rs index 7a7b4cb5..d5fdb4db 100644 --- a/rewriter/src/optimization.rs +++ b/rewriter/src/optimization.rs @@ -17,13 +17,17 @@ define_language! { "le" = Le([Id; 2]), "neg" = Neg(Id), "sqrt" = Sqrt(Id), + "log" = Log(Id), + "exp" = Exp(Id), + "xexp" = XExp(Id), + "entr" = Entr(Id), "add" = Add([Id; 2]), "sub" = Sub([Id; 2]), "mul" = Mul([Id; 2]), "div" = Div([Id; 2]), "pow" = Pow([Id; 2]), - "log" = Log(Id), - "exp" = Exp(Id), + "qol" = QOL([Id; 2]), + "norm2" = Norm2([Id; 2]), "var" = Var(Id), "param" = Param(Id), Constant(Constant), diff --git a/rewriter/src/rules.rs b/rewriter/src/rules.rs index f8ff596f..199476de 100644 --- a/rewriter/src/rules.rs +++ b/rewriter/src/rules.rs @@ -190,6 +190,17 @@ pub fn rules() -> Vec> { vec![ rw!("exp_log"; "(exp (log ?a))" => "?a" if is_gt_zero("?a")), rw!("log_exp"; "(log (exp ?a))" => "?a"), + + + /* Atom folding rules. */ + + rw!("xexp_folding"; "(mul ?a (exp ?a))" => "(xexp ?a)"), + + rw!("entr_folding"; "(neg (mul ?a (log ?a)))" => "(entr ?a)"), + + rw!("qol_folding"; "(div (pow ?a 2) ?b)" => "(qol ?a ? b)"), + + rw!("norm2"; "(sqrt (add (pow ?a 2) (pow ?b 2)))" => "(norm2 ?a ?b)"), ] } #[allow(unused)] From c357af3e6329ec9718d50b5dd226a5b4327806fa Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 09:23:23 -0500 Subject: [PATCH 018/189] fix: extra space in rule --- rewriter/src/rules.rs | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/rewriter/src/rules.rs b/rewriter/src/rules.rs index 199476de..0ff8d12f 100644 --- a/rewriter/src/rules.rs +++ b/rewriter/src/rules.rs @@ -198,7 +198,7 @@ pub fn rules() -> Vec> { vec![ rw!("entr_folding"; "(neg (mul ?a (log ?a)))" => "(entr ?a)"), - rw!("qol_folding"; "(div (pow ?a 2) ?b)" => "(qol ?a ? b)"), + rw!("qol_folding"; "(div (pow ?a 2) ?b)" => "(qol ?a ?b)"), rw!("norm2"; "(sqrt (add (pow ?a 2) (pow ?b 2)))" => "(norm2 ?a ?b)"), ] } From 79805c252b7c7b8dbcf214c72bbc0542d94eda00 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 09:23:45 -0500 Subject: [PATCH 019/189] wip: remove ad-hoc cost for composed atoms --- rewriter/src/cost.rs | 185 ++++++++++++++----------------------------- 1 file changed, 58 insertions(+), 127 deletions(-) diff --git a/rewriter/src/cost.rs b/rewriter/src/cost.rs index 91322a31..cf8594f8 100644 --- a/rewriter/src/cost.rs +++ b/rewriter/src/cost.rs @@ -15,19 +15,10 @@ pub struct DCPCost<'a> { pub egraph: &'a optimization::EGraph, } -#[derive(Debug, Clone, PartialEq)] -pub struct OptimizationWrapper(Optimization); - -impl PartialOrd for OptimizationWrapper { - fn partial_cmp(&self, other: &Self) -> Option { - Some(Ordering::Equal) - } -} - impl<'a> CostFunction for DCPCost<'a> { // Curvature + number of variables (with repetition) + term size. // In lexicographic order. - type Cost = (Curvature, u32, u32, OptimizationWrapper); + type Cost = (Curvature, u32, u32); fn cost(&mut self, enode: &Optimization, mut costs: C) -> Self::Cost where C: FnMut(Id) -> Self::Cost @@ -41,9 +32,6 @@ impl<'a> CostFunction for DCPCost<'a> { macro_rules! get_term_size { ($i:expr) => { costs(*$i).2 } } - macro_rules! get_node { - ($i:expr) => { costs(*$i).3 } - } let get_domain = |i: &Id| self.egraph[*i].data.domain.clone(); @@ -125,59 +113,29 @@ impl<'a> CostFunction for DCPCost<'a> { term_size = 1 + get_term_size!(a); } Optimization::Sqrt(a) => { - // TODO: Temporary. geo_mean. - let mut is_geo_mean = false; - match get_node!(a).0 { - Optimization::Mul([b, c]) => { - let affine = - get_curvature!(&b) == Curvature::Affine && - get_curvature!(&c) == Curvature::Affine; - let pos = - domain::option_is_pos(get_domain(&b).as_ref()) && - domain::option_is_pos(get_domain(&c).as_ref()); - is_geo_mean = affine && pos; - } - _ => {} - } - // TODO: Temporary. norm2. - let mut is_norm2 = false; - match get_node!(a).0 { - Optimization::Add([b, c]) => { - match (get_node!(&b).0, get_node!(&c).0) { - (Optimization::Pow([d, e]), Optimization::Pow([f, g])) => { - let curvature_check = - get_curvature!(&d) <= Curvature::Convex && - get_curvature!(&f) <= Curvature::Convex; - let pow_two_first = - match get_domain(&e) { - Some(de) => - match domain::Domain::get_constant(&de) { - Some (de_f) => de_f == 2.0, - _ => false - } - _ => false - }; - let pow_two_second = - match get_domain(&g) { - Some(dg) => - match domain::Domain::get_constant(&dg) { - Some (dg_f) => dg_f == 2.0, - _ => false - } - _ => false - }; - is_norm2 = curvature_check && pow_two_first && pow_two_second; - } - _ => {} - } - } - _ => {} - } - if is_geo_mean || is_norm2 { - curvature = Curvature::Convex; - } else { - curvature = curvature::of_concave_increasing_fn(get_curvature!(a)); - } + curvature = curvature::of_concave_increasing_fn(get_curvature!(a)); + num_vars = get_num_vars!(a); + term_size = 1 + get_term_size!(a); + } + Optimization::Log(a) => { + curvature = curvature::of_concave_increasing_fn(get_curvature!(a)); + num_vars = get_num_vars!(a); + term_size = 1 + get_term_size!(a); + } + Optimization::Exp(a) => { + curvature = curvature::of_convex_increasing_fn(get_curvature!(a)); + num_vars = get_num_vars!(a); + term_size = 1 + get_term_size!(a); + } + Optimization::XExp(a) => { + // TODO + curvature = curvature::of_convex_increasing_fn(get_curvature!(a)); + num_vars = get_num_vars!(a); + term_size = 1 + get_term_size!(a); + } + Optimization::Entr(a) => { + // TODO + curvature = curvature::of_concave_increasing_fn(get_curvature!(a)); num_vars = get_num_vars!(a); term_size = 1 + get_term_size!(a); } @@ -220,64 +178,35 @@ impl<'a> CostFunction for DCPCost<'a> { term_size = 1 + get_term_size!(a) + get_term_size!(b); } Optimization::Div([a, b]) => { - // TODO: Temporary. quad_over_lin. - let mut is_quad_over_lin = false; - match get_node!(a).0 { - Optimization::Pow([c, d]) => { - let curvature_check = - get_curvature!(b) <= Curvature::Concave && - get_curvature!(&c) <= Curvature::Affine && - get_curvature!(&d) == Curvature::Constant; - let pos_check = - domain::option_is_pos(get_domain(&b).as_ref()); - let pow_two_check = - match get_domain(&d) { - Some(dd) => - match domain::Domain::get_constant(&dd) { - Some (dd_f) => dd_f == 2.0, - _ => false - } - _ => false - }; - is_quad_over_lin = curvature_check && pos_check && pow_two_check; - } - _ => {} - } - - if is_quad_over_lin { - println!("quad_over_lin"); - curvature = Curvature::Convex; - } else { - let db_o = get_domain(b); - curvature = match (get_is_constant(a), get_is_constant(b)) { - (true, true) => { - match db_o { - Some(db) => { - if domain::does_not_contain_zero(&db) { - Curvature::Constant - } else { - Curvature::Unknown - } - + let db_o = get_domain(b); + curvature = match (get_is_constant(a), get_is_constant(b)) { + (true, true) => { + match db_o { + Some(db) => { + if domain::does_not_contain_zero(&db) { + Curvature::Constant + } else { + Curvature::Unknown } - None => { Curvature::Unknown } + } + None => { Curvature::Unknown } } - (false, true) => { - match db_o { - Some(db) => { - if domain::does_not_contain_zero(&db) { - curvature::of_mul_by_const(get_curvature!(a), db) - } else { - Curvature::Unknown - } + } + (false, true) => { + match db_o { + Some(db) => { + if domain::does_not_contain_zero(&db) { + curvature::of_mul_by_const(get_curvature!(a), db) + } else { + Curvature::Unknown } - None => { Curvature::Unknown } } + None => { Curvature::Unknown } } - _ => { Curvature::Unknown } - }; - } + } + _ => { Curvature::Unknown } + }; num_vars = get_num_vars!(a) + get_num_vars!(b); term_size = 1 + get_term_size!(a) + get_term_size!(b); } @@ -293,15 +222,17 @@ impl<'a> CostFunction for DCPCost<'a> { num_vars = get_num_vars!(a) + get_num_vars!(b); term_size = 1 + get_term_size!(a) + get_term_size!(b); } - Optimization::Log(a) => { - curvature = curvature::of_concave_increasing_fn(get_curvature!(a)); - num_vars = get_num_vars!(a); - term_size = 1 + get_term_size!(a); + Optimization::QOL([a, b]) => { + // TODO + curvature = Curvature::Affine; + num_vars = get_num_vars!(a) + get_num_vars!(b); + term_size = 1 + get_term_size!(a) + get_term_size!(b); } - Optimization::Exp(a) => { - curvature = curvature::of_convex_increasing_fn(get_curvature!(a)); - num_vars = get_num_vars!(a); - term_size = 1 + get_term_size!(a); + Optimization::Norm2([a, b]) => { + // TODO + curvature = Curvature::Affine; + num_vars = get_num_vars!(a) + get_num_vars!(b); + term_size = 1 + get_term_size!(a) + get_term_size!(b); } Optimization::Var(_a) => { curvature = Curvature::Affine; @@ -325,7 +256,7 @@ impl<'a> CostFunction for DCPCost<'a> { } } - return (curvature, num_vars, term_size, OptimizationWrapper(enode.clone())); + return (curvature, num_vars, term_size); } } From 2543f0afe59c54f0b85def3a26bf23eb1a6fc12a Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 09:43:31 -0500 Subject: [PATCH 020/189] fix: conditions in atom folding rules --- rewriter/src/rules.rs | 9 ++++++--- 1 file changed, 6 insertions(+), 3 deletions(-) diff --git a/rewriter/src/rules.rs b/rewriter/src/rules.rs index 0ff8d12f..9ead88f2 100644 --- a/rewriter/src/rules.rs +++ b/rewriter/src/rules.rs @@ -194,11 +194,14 @@ pub fn rules() -> Vec> { vec![ /* Atom folding rules. */ - rw!("xexp_folding"; "(mul ?a (exp ?a))" => "(xexp ?a)"), + rw!("xexp_folding"; "(mul ?a (exp ?a))" => "(xexp ?a)" + if is_ge_zero("?a")), - rw!("entr_folding"; "(neg (mul ?a (log ?a)))" => "(entr ?a)"), + rw!("entr_folding"; "(neg (mul ?a (log ?a)))" => "(entr ?a)" + if is_gt_zero("?a")), - rw!("qol_folding"; "(div (pow ?a 2) ?b)" => "(qol ?a ?b)"), + rw!("qol_folding"; "(div (pow ?a 2) ?b)" => "(qol ?a ?b)" + if is_gt_zero("?b")), rw!("norm2"; "(sqrt (add (pow ?a 2) (pow ?b 2)))" => "(norm2 ?a ?b)"), ] } From b2e11601c1213292e4344b9d21ed70ae9f62050b Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 09:55:06 -0500 Subject: [PATCH 021/189] feat: all composition rules and joining curvatures --- rewriter/src/curvature.rs | 51 +++++++++++++++++++++++++++++++++++++-- 1 file changed, 49 insertions(+), 2 deletions(-) diff --git a/rewriter/src/curvature.rs b/rewriter/src/curvature.rs index 3708bc96..45cdc1ed 100644 --- a/rewriter/src/curvature.rs +++ b/rewriter/src/curvature.rs @@ -78,6 +78,10 @@ impl fmt::Display for Curvature { } } +pub fn join(c1: Curvature, c2: Curvature) -> Curvature { + if c1 <= c2 { c2 } else if c2 <= c1 { c1 } else { Unknown } +} + pub fn of_le(c1: Curvature, c2: Curvature) -> Curvature { match (c1, c2) { (Convex, Concave ) => { return Convex; } @@ -103,7 +107,8 @@ pub fn of_neg(c: Curvature) -> Curvature { } } -pub fn of_concave_increasing_fn(c: Curvature) -> Curvature { +// Composition rule for nondecreasing (or increasing) concave functions. +pub fn of_concave_nondecreasing_fn(c: Curvature) -> Curvature { match c { Concave => { return Concave; } Affine => { return Concave; } @@ -112,7 +117,28 @@ pub fn of_concave_increasing_fn(c: Curvature) -> Curvature { } } -pub fn of_convex_increasing_fn(c: Curvature) -> Curvature { +// Composition rule for nonincreasing (or decreasing) concave functions. +#[allow(unused)] +pub fn of_concave_nonincreasing_fn(c: Curvature) -> Curvature { + match c { + Convex => { return Concave; } + Affine => { return Concave; } + Constant => { return Constant; } + _ => { return Unknown; } + } +} + +// Composition rule for concave functions with unknown monotonocity. +pub fn of_concave_none_fn(c: Curvature) -> Curvature { + match c { + Affine => { return Concave; } + Constant => { return Constant; } + _ => { return Unknown; } + } +} + +// Composition rule for nondecreasing (or increasing) convex functions. +pub fn of_convex_nondecreasing_fn(c: Curvature) -> Curvature { match c { Convex => { return Convex; } Affine => { return Convex; } @@ -121,6 +147,27 @@ pub fn of_convex_increasing_fn(c: Curvature) -> Curvature { } } +// Composition rule for nonincreasing (or decreasing) convex functions. +#[allow(unused)] +pub fn of_convex_nonincreasing_fn(c: Curvature) -> Curvature { + match c { + Concave => { return Convex; } + Affine => { return Convex; } + Constant => { return Constant; } + _ => { return Unknown; } + } +} + +// Composition rule for concave functions with unknown monotonocity. +#[allow(unused)] +pub fn of_convex_none_fn(c: Curvature) -> Curvature { + match c { + Affine => { return Convex; } + Constant => { return Constant; } + _ => { return Unknown; } + } +} + pub fn of_add(c1: Curvature, c2: Curvature) -> Curvature { match (c1, c2) { (Convex, Convex ) => { return Convex; } From 625bf9831d0e0a2f7c18746c84835127ef8748ab Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 09:55:33 -0500 Subject: [PATCH 022/189] feat: cost calculation for new atoms --- rewriter/src/cost.rs | 42 +++++++++++++++++++++++++++++------------- 1 file changed, 29 insertions(+), 13 deletions(-) diff --git a/rewriter/src/cost.rs b/rewriter/src/cost.rs index cf8594f8..a71e504f 100644 --- a/rewriter/src/cost.rs +++ b/rewriter/src/cost.rs @@ -1,5 +1,3 @@ -use std::cmp::Ordering; - use egg::{*}; use crate::domain; @@ -113,29 +111,27 @@ impl<'a> CostFunction for DCPCost<'a> { term_size = 1 + get_term_size!(a); } Optimization::Sqrt(a) => { - curvature = curvature::of_concave_increasing_fn(get_curvature!(a)); + curvature = curvature::of_concave_nondecreasing_fn(get_curvature!(a)); num_vars = get_num_vars!(a); term_size = 1 + get_term_size!(a); } Optimization::Log(a) => { - curvature = curvature::of_concave_increasing_fn(get_curvature!(a)); + curvature = curvature::of_concave_nondecreasing_fn(get_curvature!(a)); num_vars = get_num_vars!(a); term_size = 1 + get_term_size!(a); } Optimization::Exp(a) => { - curvature = curvature::of_convex_increasing_fn(get_curvature!(a)); + curvature = curvature::of_convex_nondecreasing_fn(get_curvature!(a)); num_vars = get_num_vars!(a); term_size = 1 + get_term_size!(a); } Optimization::XExp(a) => { - // TODO - curvature = curvature::of_convex_increasing_fn(get_curvature!(a)); + curvature = curvature::of_convex_nondecreasing_fn(get_curvature!(a)); num_vars = get_num_vars!(a); term_size = 1 + get_term_size!(a); } Optimization::Entr(a) => { - // TODO - curvature = curvature::of_concave_increasing_fn(get_curvature!(a)); + curvature = curvature::of_concave_none_fn(get_curvature!(a)); num_vars = get_num_vars!(a); term_size = 1 + get_term_size!(a); } @@ -223,14 +219,34 @@ impl<'a> CostFunction for DCPCost<'a> { term_size = 1 + get_term_size!(a) + get_term_size!(b); } Optimization::QOL([a, b]) => { - // TODO - curvature = Curvature::Affine; + let da_nonneg = domain::option_is_nonneg(get_domain(a).as_ref()); + let curvature_num = + if da_nonneg { + curvature::of_convex_nondecreasing_fn(get_curvature!(a)) + } else { + curvature::of_convex_nonincreasing_fn(get_curvature!(a)) + }; + let curvature_den = curvature::of_convex_nonincreasing_fn(get_curvature!(b)); + curvature = curvature::join(curvature_num, curvature_den); num_vars = get_num_vars!(a) + get_num_vars!(b); term_size = 1 + get_term_size!(a) + get_term_size!(b); } Optimization::Norm2([a, b]) => { - // TODO - curvature = Curvature::Affine; + let da_nonneg = domain::option_is_nonneg(get_domain(a).as_ref()); + let curvature_a = + if da_nonneg { + curvature::of_convex_nondecreasing_fn(get_curvature!(a)) + } else { + curvature::of_convex_nonincreasing_fn(get_curvature!(a)) + }; + let db_nonneg = domain::option_is_nonneg(get_domain(b).as_ref()); + let curvature_b = + if db_nonneg { + curvature::of_convex_nondecreasing_fn(get_curvature!(b)) + } else { + curvature::of_convex_nonincreasing_fn(get_curvature!(b)) + }; + curvature = curvature::join(curvature_a, curvature_b); num_vars = get_num_vars!(a) + get_num_vars!(b); term_size = 1 + get_term_size!(a) + get_term_size!(b); } From 606b13689582808c9755f51d21180c242e98e311 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 10:03:54 -0500 Subject: [PATCH 023/189] feat: add domain calculation for new atoms --- rewriter/src/domain.rs | 60 ++++++++++++++++++++++++++++++------ rewriter/src/optimization.rs | 33 ++++++++++++++++---- 2 files changed, 77 insertions(+), 16 deletions(-) diff --git a/rewriter/src/domain.rs b/rewriter/src/domain.rs index d9029284..1b9f6ff7 100644 --- a/rewriter/src/domain.rs +++ b/rewriter/src/domain.rs @@ -482,6 +482,22 @@ pub fn option_exp(d_o: Option) -> Option { } } +pub fn xexp(d: &Domain) -> Domain { + mul(d, &exp(d)) +} + +pub fn option_xexp(d_o: Option) -> Option { + execute_unary(d_o, xexp) +} + +pub fn entr(d: &Domain) -> Domain { + neg(&mul(d, &log(d))) +} + +pub fn option_entr(d_o: Option) -> Option { + execute_unary(d_o, entr) +} + pub fn add(d1: &Domain, d2: &Domain) -> Domain { let l1 = d1.lo_open; let r1 = d1.hi_open; @@ -494,8 +510,8 @@ pub fn add(d1: &Domain, d2: &Domain) -> Domain { ) } -pub fn option_add(d_o_a: Option, d_o_b: Option) -> Option { - execute_binary(d_o_a, d_o_b, add) +pub fn option_add(d1_o: Option, d2_o: Option) -> Option { + execute_binary(d1_o, d2_o, add) } pub fn sub(d1: &Domain, d2: &Domain) -> Domain { @@ -510,8 +526,8 @@ pub fn sub(d1: &Domain, d2: &Domain) -> Domain { ) } -pub fn option_sub(d_o_a: Option, d_o_b: Option) -> Option { - execute_binary(d_o_a, d_o_b, sub) +pub fn option_sub(d1_o: Option, d2_o: Option) -> Option { + execute_binary(d1_o, d2_o, sub) } // The idea is that open "beats" closed. @@ -617,8 +633,8 @@ pub fn mul(d1: &Domain, d2: &Domain) -> Domain { } -pub fn option_mul(d_o_a: Option, d_o_b: Option) -> Option { - execute_binary(d_o_a, d_o_b, mul) +pub fn option_mul(d1_o: Option, d2_o: Option) -> Option { + execute_binary(d1_o, d2_o, mul) } // Copied from intervals-good. Again, zero divided by anything is zero, with the @@ -702,8 +718,8 @@ pub fn div(d1: &Domain, d2: &Domain) -> Domain { } } -pub fn option_div(d_o_a: Option, d_o_b: Option) -> Option { - execute_binary(d_o_a, d_o_b, div) +pub fn option_div(d1_o: Option, d2_o: Option) -> Option { + execute_binary(d1_o, d2_o, div) } // Same reasoning as in `perform_pow` @@ -809,6 +825,30 @@ pub fn pow(d1: &Domain, d2: &Domain) -> Domain { } } -pub fn option_pow(d_o_a: Option, d_o_b: Option) -> Option { - execute_binary(d_o_a, d_o_b, pow) +pub fn option_pow(d1_o: Option, d2_o: Option) -> Option { + execute_binary(d1_o, d2_o, pow) +} + +pub fn quad_over_lin(d1: &Domain, d2: &Domain) -> Domain { + div(&pow(d1, &Domain::make_singleton(2.0)), d2) +} + +pub fn option_quad_over_lin(d1_o: Option, d2_o: Option) -> Option { + execute_binary(d1_o, d2_o, quad_over_lin) +} + +pub fn geo_mean(d1: &Domain, d2: &Domain) -> Domain { + sqrt(&mul(d1, d2)) +} + +pub fn option_geo_mean(d1_o: Option, d2_o: Option) -> Option { + execute_binary(d1_o, d2_o, geo_mean) +} + +pub fn norm2(d1: &Domain, d2: &Domain) -> Domain { + sqrt(&add(&pow(d1, &Domain::make_singleton(2.0)), &pow(d2, &Domain::make_singleton(2.0)))) +} + +pub fn option_norm2(d1_o: Option, d2_o: Option) -> Option { + execute_binary(d1_o, d2_o, norm2) } diff --git a/rewriter/src/optimization.rs b/rewriter/src/optimization.rs index d5fdb4db..00a559eb 100644 --- a/rewriter/src/optimization.rs +++ b/rewriter/src/optimization.rs @@ -27,6 +27,7 @@ define_language! { "div" = Div([Id; 2]), "pow" = Pow([Id; 2]), "qol" = QOL([Id; 2]), + "geo" = Geo([Id; 2]), "norm2" = Norm2([Id; 2]), "var" = Var(Id), "param" = Param(Id), @@ -112,6 +113,22 @@ impl Analysis for Meta { domain = domain::option_sqrt(get_domain(a)); is_constant = get_is_constant(a); } + Optimization::Log(a) => { + domain = domain::option_log(get_domain(a)); + is_constant = get_is_constant(a); + } + Optimization::Exp(a) => { + domain = domain::option_exp(get_domain(a)); + is_constant = get_is_constant(a); + } + Optimization::XExp(a) => { + domain = domain::option_xexp(get_domain(a)); + is_constant = get_is_constant(a); + } + Optimization::Entr(a) => { + domain = domain::option_entr(get_domain(a)); + is_constant = get_is_constant(a); + } Optimization::Add([a, b]) => { domain = domain::option_add( get_domain(a), get_domain(b)); @@ -133,13 +150,17 @@ impl Analysis for Meta { domain = domain::option_pow(get_domain(a), get_domain(b)); is_constant = get_is_constant(a) && get_is_constant(b); } - Optimization::Log(a) => { - domain = domain::option_log(get_domain(a)); - is_constant = get_is_constant(a); + Optimization::QOL([a, b]) => { + domain = domain::option_quad_over_lin(get_domain(a), get_domain(b)); + is_constant = get_is_constant(a) && get_is_constant(b); } - Optimization::Exp(a) => { - domain = domain::option_exp(get_domain(a)); - is_constant = get_is_constant(a); + Optimization::Geo([a, b]) => { + domain = domain::option_geo_mean(get_domain(a), get_domain(b)); + is_constant = get_is_constant(a) && get_is_constant(b); + } + Optimization::Norm2([a, b]) => { + domain = domain::option_norm2(get_domain(a), get_domain(b)); + is_constant = get_is_constant(a) && get_is_constant(b); } Optimization::Var(a) => { // Assume that after var there is always a symbol. From 78fef5cc8a766a89197ece0523ce648fe4bcbf42 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 10:06:41 -0500 Subject: [PATCH 024/189] feat: cost and rule for `geo_mean` --- rewriter/src/cost.rs | 7 +++++++ rewriter/src/rules.rs | 5 ++++- 2 files changed, 11 insertions(+), 1 deletion(-) diff --git a/rewriter/src/cost.rs b/rewriter/src/cost.rs index a71e504f..19f8aae0 100644 --- a/rewriter/src/cost.rs +++ b/rewriter/src/cost.rs @@ -231,6 +231,13 @@ impl<'a> CostFunction for DCPCost<'a> { num_vars = get_num_vars!(a) + get_num_vars!(b); term_size = 1 + get_term_size!(a) + get_term_size!(b); } + Optimization::Geo([a, b]) => { + let curvature_a = curvature::of_convex_nondecreasing_fn(get_curvature!(a)); + let curvature_b = curvature::of_convex_nondecreasing_fn(get_curvature!(b)); + curvature = curvature::join(curvature_a, curvature_b); + num_vars = get_num_vars!(a) + get_num_vars!(b); + term_size = 1 + get_term_size!(a) + get_term_size!(b); + } Optimization::Norm2([a, b]) => { let da_nonneg = domain::option_is_nonneg(get_domain(a).as_ref()); let curvature_a = diff --git a/rewriter/src/rules.rs b/rewriter/src/rules.rs index 9ead88f2..9f983fae 100644 --- a/rewriter/src/rules.rs +++ b/rewriter/src/rules.rs @@ -203,7 +203,10 @@ pub fn rules() -> Vec> { vec![ rw!("qol_folding"; "(div (pow ?a 2) ?b)" => "(qol ?a ?b)" if is_gt_zero("?b")), - rw!("norm2"; "(sqrt (add (pow ?a 2) (pow ?b 2)))" => "(norm2 ?a ?b)"), + rw!("geo_folding"; "(sqrt (mul ?a ?b))" => "(geo ?a ?b)" + if is_gt_zero("?a") if is_gt_zero("?b")), + + rw!("norm2_folding"; "(sqrt (add (pow ?a 2) (pow ?b 2)))" => "(norm2 ?a ?b)"), ] } #[allow(unused)] From ca8252954b4ace92dd0296a47589c63ef82cb740 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 10:29:29 -0500 Subject: [PATCH 025/189] feat: `get_steps_from_string` --- rewriter/src/extract.rs | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/rewriter/src/extract.rs b/rewriter/src/extract.rs index 36c1ceae..9446066e 100644 --- a/rewriter/src/extract.rs +++ b/rewriter/src/extract.rs @@ -148,12 +148,13 @@ impl ToString for Minimization { } } -/// Return the rewrite steps if egg successfully found a chain of rewrites to -/// transform the term into DCP form. Return `None` if it didn't. +// Return the rewrite steps if egg successfully found a chain of rewrites to +// transform the term into DCP form. Return `None` if it didn't. pub fn get_steps(prob: Minimization, domains_vec: Vec<(String, Domain)>, debug: bool) -> Option> { - let prob_s = prob.to_string(); - let expr: RecExpr = prob_s.parse().unwrap(); + get_steps_from_string(&prob.to_string(), domains_vec, debug) +} +pub fn get_steps_from_string(prob_s: &str, domains_vec: Vec<(String, Domain)>, debug: bool) -> Option> { // Process domains, intersecting domains assigned to the same variable. let mut domains: HashMap = HashMap::new(); for (x, dom) in domains_vec { From fd72a6a18632786ccbdb4d593fb9d375f62e0de5 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 10:34:57 -0500 Subject: [PATCH 026/189] fix: missing variable in extractor --- rewriter/src/extract.rs | 2 ++ 1 file changed, 2 insertions(+) diff --git a/rewriter/src/extract.rs b/rewriter/src/extract.rs index 9446066e..110324f7 100644 --- a/rewriter/src/extract.rs +++ b/rewriter/src/extract.rs @@ -155,6 +155,8 @@ pub fn get_steps(prob: Minimization, domains_vec: Vec<(String, Domain)>, debug: } pub fn get_steps_from_string(prob_s: &str, domains_vec: Vec<(String, Domain)>, debug: bool) -> Option> { + let expr: RecExpr = prob_s.parse().unwrap(); + // Process domains, intersecting domains assigned to the same variable. let mut domains: HashMap = HashMap::new(); for (x, dom) in domains_vec { From 2af01a01920e48eb9bec191536163bd44c52310e Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 10:35:19 -0500 Subject: [PATCH 027/189] wip: restricted constant folding --- rewriter/src/optimization.rs | 25 +++++++++++++++++++++++++ 1 file changed, 25 insertions(+) diff --git a/rewriter/src/optimization.rs b/rewriter/src/optimization.rs index 00a559eb..8b63cc75 100644 --- a/rewriter/src/optimization.rs +++ b/rewriter/src/optimization.rs @@ -199,6 +199,31 @@ impl Analysis for Meta { Data { domain, is_constant } } + + fn modify(egraph: &mut egg::EGraph, id: Id) { + let data = egraph[id].data.clone(); + if data.is_constant { + match data.domain { + Some(d) => { + match d.get_constant() { + Some(c) => { + // Only fold if constant is .0 or .5. + if ((2.0 * c) as u32) as f64 == (2.0 * c) { + let nn_c = NotNan::new(c).unwrap(); + let node = Optimization::Constant(nn_c); + let added = egraph.add(node); + egraph.union(id, added); + + egraph[id].assert_unique_leaves(); + } + } + _ => {} + } + } + _ => {} + } + } + } } pub fn is_gt_zero(var: &str) -> impl Fn(&mut EGraph, Id, &Subst) -> bool { From 8fc1b076b8d2a41a10ecbf0d9e8d2decf6021df1 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 10:36:26 -0500 Subject: [PATCH 028/189] feat: tests outside of optimization problems --- rewriter/src/tests.rs | 26 ++++++++++++++++++++------ 1 file changed, 20 insertions(+), 6 deletions(-) diff --git a/rewriter/src/tests.rs b/rewriter/src/tests.rs index ba5f11e7..7149616b 100644 --- a/rewriter/src/tests.rs +++ b/rewriter/src/tests.rs @@ -1,12 +1,13 @@ #[cfg(test)] mod tests { -use crate::domain; +use crate::domain; use domain::Domain as Domain; use crate::extract; use extract::Minimization as Minimization; use extract::get_steps as get_steps; +use extract::get_steps_from_string as get_steps_from_string; fn make(obj: &str, constrs: Vec<&str>) -> Minimization { let mut constrs_s = Vec::new(); @@ -33,6 +34,20 @@ fn assert_steps(obj: &str, constrs: Vec<&str>) { assert_steps_with_domain(vec![], obj, constrs); } +// Used to test single expressions outside the context of an optimization +// problem. +fn assert_steps_from_string_with_domain(domains : Vec<(&str, Domain)>, s: &str) { + let domains = + domains.iter().map(|(s, d)| ((*s).to_string(), d.clone())).collect(); + let steps = get_steps_from_string(s, domains, true); + println!("{:?}", steps); + assert!(steps.is_some()); +} + +fn assert_steps_from_string(s: &str) { + assert_steps_from_string_with_domain(vec![], s); +} + // Examples. @@ -270,11 +285,9 @@ fn test_norm2() { #[test] fn test_3_32() { - assert_steps_with_domain( + assert_steps_from_string_with_domain( vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], - "(div 1 (mul (var x) (var y)))", - vec![ - ]); + "(div 1 (mul (var x) (var y)))"); } #[test] @@ -285,7 +298,8 @@ fn test_3_33() { assert_steps_with_domain( vec![("x", domain::free_dom()), ("y", domain::pos_dom())], // "(sqrt (add 1 (div (pow (var x) 4) (var y))))", - "(sqrt (add 1 (div (pow (var x) (mul 2 2)) (var y))))", + // "(sqrt (add 1 (div (pow (var x) (mul 2 2)) (var y))))", + "(div (pow (var x) 2) (sqrt (var y)) )", vec![ ]); } From 40dba18119bdd2478f8a897e45639b6bb71c9747 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 10:40:06 -0500 Subject: [PATCH 029/189] chore: remove unused tag --- rewriter/src/curvature.rs | 1 - 1 file changed, 1 deletion(-) diff --git a/rewriter/src/curvature.rs b/rewriter/src/curvature.rs index 45cdc1ed..92f15786 100644 --- a/rewriter/src/curvature.rs +++ b/rewriter/src/curvature.rs @@ -148,7 +148,6 @@ pub fn of_convex_nondecreasing_fn(c: Curvature) -> Curvature { } // Composition rule for nonincreasing (or decreasing) convex functions. -#[allow(unused)] pub fn of_convex_nonincreasing_fn(c: Curvature) -> Curvature { match c { Concave => { return Convex; } From a4a45742b1c1c22aafe0bfae17826de98d93513c Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 10:48:59 -0500 Subject: [PATCH 030/189] chore: increase node limit --- rewriter/src/extract.rs | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/rewriter/src/extract.rs b/rewriter/src/extract.rs index 110324f7..2fa47d15 100644 --- a/rewriter/src/extract.rs +++ b/rewriter/src/extract.rs @@ -172,7 +172,7 @@ pub fn get_steps_from_string(prob_s: &str, domains_vec: Vec<(String, Domain)>, d } } - for node_limit in [2500, 5000, 7500, 10000] { + for node_limit in [2500, 5000, 10000, 20000] { let analysis = Meta { domains : domains.clone() }; From 9235fae1c180bbf2e2963217b68429b0773b03c0 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 10:49:16 -0500 Subject: [PATCH 031/189] fix: geometric mean is concave --- rewriter/src/cost.rs | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/rewriter/src/cost.rs b/rewriter/src/cost.rs index 19f8aae0..9ab669af 100644 --- a/rewriter/src/cost.rs +++ b/rewriter/src/cost.rs @@ -232,8 +232,8 @@ impl<'a> CostFunction for DCPCost<'a> { term_size = 1 + get_term_size!(a) + get_term_size!(b); } Optimization::Geo([a, b]) => { - let curvature_a = curvature::of_convex_nondecreasing_fn(get_curvature!(a)); - let curvature_b = curvature::of_convex_nondecreasing_fn(get_curvature!(b)); + let curvature_a = curvature::of_concave_nondecreasing_fn(get_curvature!(a)); + let curvature_b = curvature::of_concave_nondecreasing_fn(get_curvature!(b)); curvature = curvature::join(curvature_a, curvature_b); num_vars = get_num_vars!(a) + get_num_vars!(b); term_size = 1 + get_term_size!(a) + get_term_size!(b); From e52041c248594d12ab0658add7ce7eb14a79647d Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 10:58:44 -0500 Subject: [PATCH 032/189] feat: check if any domain is empty --- rewriter/src/extract.rs | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/rewriter/src/extract.rs b/rewriter/src/extract.rs index 2fa47d15..8c023e12 100644 --- a/rewriter/src/extract.rs +++ b/rewriter/src/extract.rs @@ -171,6 +171,12 @@ pub fn get_steps_from_string(prob_s: &str, domains_vec: Vec<(String, Domain)>, d domains.insert(x, dom); } } + for (_, d) in domains.clone() { + if d.is_empty() { + println!("Unfeasible problem."); + return None; + } + } for node_limit in [2500, 5000, 10000, 20000] { let analysis = Meta { From a686eeb603ae21bed0ebeb79a8f6c8af38e400fa Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 11:15:45 -0500 Subject: [PATCH 033/189] feat: simplify extractor and allow out-of-context expressions --- rewriter/src/explain_utils.rs | 4 ++++ rewriter/src/extract.rs | 31 ++++++++----------------------- 2 files changed, 12 insertions(+), 23 deletions(-) diff --git a/rewriter/src/explain_utils.rs b/rewriter/src/explain_utils.rs index 42924605..c83b3b41 100644 --- a/rewriter/src/explain_utils.rs +++ b/rewriter/src/explain_utils.rs @@ -39,6 +39,7 @@ pub fn make_rule_table(rules: &Vec) -> HashMap { // These are functions taken from `explain.rs` removing the assertions. +#[allow(unused)] fn flat_term_make_bindings_no_failure<'a>( term: &'a FlatTerm, pattern: &[ENodeOrVar], @@ -70,6 +71,7 @@ fn flat_term_make_bindings_no_failure<'a>( } } +#[allow(unused)] fn flat_term_from_pattern( pattern: &[ENodeOrVar], location: usize, @@ -91,6 +93,7 @@ fn flat_term_from_pattern( } } +#[allow(unused)] fn flat_term_rewrite_no_failure( term: &FlatTerm, lhs: &PatternAst, @@ -105,6 +108,7 @@ fn flat_term_rewrite_no_failure( return Some(flat_term_from_pattern(rhs_nodes, rhs_nodes.len() - 1, &bindings)); } +#[allow(unused)] pub fn flat_term_check_rewrite_no_failure<'a>( current: &'a FlatTerm, next: &'a FlatTerm, diff --git a/rewriter/src/extract.rs b/rewriter/src/extract.rs index 8c023e12..0e9de70f 100644 --- a/rewriter/src/extract.rs +++ b/rewriter/src/extract.rs @@ -24,7 +24,6 @@ use crate::explain_utils; use explain_utils::Direction as Direction; use explain_utils::make_rule_table as make_rule_table; use explain_utils::get_rewrite_name_and_direction as get_rewrite_name_and_direction; -use explain_utils::flat_term_check_rewrite_no_failure as flat_term_check_rewrite_no_failure; pub type Rewrite = egg::Rewrite; @@ -33,7 +32,6 @@ pub struct Step { rewrite_name : String, direction : Direction, location : String, - // position : u32, expected_term : String, } @@ -42,8 +40,7 @@ fn get_step_aux( direction: Direction, current: &FlatTerm, next: &FlatTerm, - location: &mut Option, - position : &mut u32, + location: &mut Option, expected_term: &mut Option) -> Option { match next.node { @@ -63,22 +60,12 @@ fn get_step_aux( _ => () } - // Check if it matches to update the position. - let matches = match direction { - Direction::Forward => flat_term_check_rewrite_no_failure(current, next, rewrite), - Direction::Backward => flat_term_check_rewrite_no_failure(next, current, rewrite), - }; - if matches { - *position += 1; - } - if let Some(rule_name) = &next.backward_rule { - if location.is_some() && expected_term.is_some() { + if expected_term.is_some() { return Some(Step { rewrite_name: rule_name.to_string(), direction: Direction::Backward, - location: location.clone().unwrap(), - // position: position.clone(), + location: location.clone().unwrap_or_default(), expected_term: expected_term.clone().unwrap(), }); } else { @@ -91,12 +78,11 @@ fn get_step_aux( // let next_s_t = next.get_recexpr().to_string(); // println!("curr {} : {}", rule_name, curr_s_t); // println!("next {} : {}", rule_name, next_s_t); - if location.is_some() && expected_term.is_some() { + if expected_term.is_some() { return Some(Step { rewrite_name: rule_name.to_string(), direction: Direction::Forward, - location: location.clone().unwrap(), - // position: position.clone(), + location: location.clone().unwrap_or_default(), expected_term: expected_term.clone().unwrap(), }); } else { @@ -110,7 +96,7 @@ fn get_step_aux( let children = current.children.iter().zip(next.children.iter()); for (left, right) in children { let child_res = - get_step_aux(rewrite, direction, left, right, location, position, expected_term); + get_step_aux(rewrite, direction, left, right, location, expected_term); if child_res.is_some() { return child_res; } @@ -124,9 +110,8 @@ fn get_step(rule_table: &HashMap, current: &FlatTerm, d let rules_copy = rules().clone(); let rule_table = make_rule_table(&rules_copy); - + let mut res = Vec::new(); if best_cost.0 <= Curvature::Convex { for i in 0..flat_explanation.len() - 1 { From da7630272cf19fda2b0445f6d576834cd48a942e Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 11:36:58 -0500 Subject: [PATCH 034/189] feat: special cases for powers with general bases --- rewriter/src/domain.rs | 31 ++++++++++++++++++++++++++++--- 1 file changed, 28 insertions(+), 3 deletions(-) diff --git a/rewriter/src/domain.rs b/rewriter/src/domain.rs index 1b9f6ff7..9aa78171 100644 --- a/rewriter/src/domain.rs +++ b/rewriter/src/domain.rs @@ -376,6 +376,10 @@ pub fn is_nonpos(d: &Domain) -> bool { d.subseteq(&nonpos_dom()) } +pub fn option_is_nonpos(d: Option<&Domain>) -> bool { + d.map_or(false, is_nonpos) +} + pub fn pos_dom() -> Domain { Domain::make(nonneg_ival(), true, false) } @@ -757,9 +761,30 @@ fn perform_pow( pub fn pow(d1: &Domain, d2: &Domain) -> Domain { let d1_nonneg = is_nonneg(d1); if !d1_nonneg { - // We only define it for nonnegative bases. Otherwise, return the - // free domain. - return free_dom(); + // We only define rules for nonnegative bases. + // However, if the exponent is a constant, we consider some special + // cases: 0, 1, and even integers. + match d2.get_constant() { + Some(f) => { + if ((f as u32) as f64) == f { + let n = f as u32; + if n == 0 && does_not_contain_zero(d1) { + return Domain::make_singleton(1.0); + } else if n == 1 { + return d1.clone(); + } else if n % 2 == 0 { + return nonneg_dom(); + } else { + return free_dom(); + } + } else { + return free_dom(); + } + } + _ => { + return free_dom(); + } + } } let d1_lrg = d1.subseteq(&Domain::make_oi(one())); From 6b6f6a084e379dfd5e7e24f80780b37f3e1d0d45 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 11:38:17 -0500 Subject: [PATCH 035/189] fix: careful handling of monotonicity ranges --- rewriter/src/cost.rs | 15 ++++++++++++--- 1 file changed, 12 insertions(+), 3 deletions(-) diff --git a/rewriter/src/cost.rs b/rewriter/src/cost.rs index 9ab669af..f3f100b0 100644 --- a/rewriter/src/cost.rs +++ b/rewriter/src/cost.rs @@ -220,11 +220,14 @@ impl<'a> CostFunction for DCPCost<'a> { } Optimization::QOL([a, b]) => { let da_nonneg = domain::option_is_nonneg(get_domain(a).as_ref()); + let da_nonpos = domain::option_is_nonpos(get_domain(a).as_ref()); let curvature_num = if da_nonneg { curvature::of_convex_nondecreasing_fn(get_curvature!(a)) - } else { + } else if da_nonpos { curvature::of_convex_nonincreasing_fn(get_curvature!(a)) + } else { + curvature::of_convex_none_fn(get_curvature!(a)) }; let curvature_den = curvature::of_convex_nonincreasing_fn(get_curvature!(b)); curvature = curvature::join(curvature_num, curvature_den); @@ -240,18 +243,24 @@ impl<'a> CostFunction for DCPCost<'a> { } Optimization::Norm2([a, b]) => { let da_nonneg = domain::option_is_nonneg(get_domain(a).as_ref()); + let da_nonpos = domain::option_is_nonpos(get_domain(a).as_ref()); let curvature_a = if da_nonneg { curvature::of_convex_nondecreasing_fn(get_curvature!(a)) - } else { + } else if da_nonpos { curvature::of_convex_nonincreasing_fn(get_curvature!(a)) + } else { + curvature::of_convex_none_fn(get_curvature!(a)) }; let db_nonneg = domain::option_is_nonneg(get_domain(b).as_ref()); + let db_nonpos = domain::option_is_nonpos(get_domain(b).as_ref()); let curvature_b = if db_nonneg { curvature::of_convex_nondecreasing_fn(get_curvature!(b)) - } else { + } else if db_nonpos { curvature::of_convex_nonincreasing_fn(get_curvature!(b)) + } else { + curvature::of_convex_none_fn(get_curvature!(b)) }; curvature = curvature::join(curvature_a, curvature_b); num_vars = get_num_vars!(a) + get_num_vars!(b); From 7f597b1f59472e9fbcae3a41c4797ad72faefa94 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 11:38:37 -0500 Subject: [PATCH 036/189] fix: extractor find the right expected term --- rewriter/src/extract.rs | 9 ++++++++- 1 file changed, 8 insertions(+), 1 deletion(-) diff --git a/rewriter/src/extract.rs b/rewriter/src/extract.rs index 0e9de70f..f4e0a0f2 100644 --- a/rewriter/src/extract.rs +++ b/rewriter/src/extract.rs @@ -44,6 +44,8 @@ fn get_step_aux( expected_term: &mut Option) -> Option { match next.node { + Optimization::Prob(_) => { }, + Optimization::Constrs(_) => { }, Optimization::ObjFun(_) => { *location = Some("objFun".to_string()); @@ -57,7 +59,11 @@ fn get_step_aux( let c_s = next.children[1].get_recexpr().to_string(); *expected_term = Some(c_s); } - _ => () + _ => { + // Out-of-context extraction. Useful for testing. + let c_s = next.get_recexpr().to_string(); + *expected_term = Some(c_s); + } } if let Some(rule_name) = &next.backward_rule { @@ -74,6 +80,7 @@ fn get_step_aux( } if let Some(rule_name) = &next.forward_rule { + // NOTE: Could be useful for more targetted rewrites. // let curr_s_t: String = current.get_recexpr().to_string(); // let next_s_t = next.get_recexpr().to_string(); // println!("curr {} : {}", rule_name, curr_s_t); From c5dd4d9bed5d08b9c7b715e0e5c26c9b2046152b Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 11:38:56 -0500 Subject: [PATCH 037/189] feat: some power rules --- rewriter/src/rules.rs | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/rewriter/src/rules.rs b/rewriter/src/rules.rs index 9f983fae..cf53f16d 100644 --- a/rewriter/src/rules.rs +++ b/rewriter/src/rules.rs @@ -111,6 +111,8 @@ pub fn rules() -> Vec> { vec![ /* Power and square root rules. */ + rw!("one_pow"; "(pow 1 ?a)" => "1"), + rw!("pow_add"; "(pow ?a (add ?b ?c))" => "(mul (pow ?a ?b) (pow ?a ?c))" if is_gt_zero("?a")), @@ -152,10 +154,10 @@ pub fn rules() -> Vec> { vec![ rw!("sqrt_eq_rpow"; "(sqrt ?a)" => "(pow ?a 0.5)"), - // NOTE: Needed since constant folding is disabled. - // TODO: Activate constant folding but only for rationals? rw!("pow_half_two"; "(pow (pow ?a 0.5) 2)" => "?a" if is_ge_zero("?a")), + rw!("pow_half_two-rev"; "?a" => "(pow (pow ?a 0.5) 2)" if is_ge_zero("?a")), + /* Exponential and logarithm rules. */ From 479ddbdc979f6ba45e6de8ab2a1cfd9d6b672c0e Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 12:14:47 -0500 Subject: [PATCH 038/189] test: 3.33 passing --- rewriter/src/domain.rs | 1 + rewriter/src/rules.rs | 4 ++++ rewriter/src/tests.rs | 51 +++++++++++++++++++++++------------------- 3 files changed, 33 insertions(+), 23 deletions(-) diff --git a/rewriter/src/domain.rs b/rewriter/src/domain.rs index 9aa78171..3569e44b 100644 --- a/rewriter/src/domain.rs +++ b/rewriter/src/domain.rs @@ -773,6 +773,7 @@ pub fn pow(d1: &Domain, d2: &Domain) -> Domain { } else if n == 1 { return d1.clone(); } else if n % 2 == 0 { + println!("here {}", n); return nonneg_dom(); } else { return free_dom(); diff --git a/rewriter/src/rules.rs b/rewriter/src/rules.rs index cf53f16d..fbe798cc 100644 --- a/rewriter/src/rules.rs +++ b/rewriter/src/rules.rs @@ -113,6 +113,8 @@ pub fn rules() -> Vec> { vec![ rw!("one_pow"; "(pow 1 ?a)" => "1"), + rw!("pow_one"; "(pow ?a 1)" => "?a"), + rw!("pow_add"; "(pow ?a (add ?b ?c))" => "(mul (pow ?a ?b) (pow ?a ?c))" if is_gt_zero("?a")), @@ -154,6 +156,8 @@ pub fn rules() -> Vec> { vec![ rw!("sqrt_eq_rpow"; "(sqrt ?a)" => "(pow ?a 0.5)"), + rw!("sqrt_eq_rpow-rev"; "(pow ?a 0.5)" => "(sqrt ?a)"), + rw!("pow_half_two"; "(pow (pow ?a 0.5) 2)" => "?a" if is_ge_zero("?a")), rw!("pow_half_two-rev"; "?a" => "(pow (pow ?a 0.5) 2)" if is_ge_zero("?a")), diff --git a/rewriter/src/tests.rs b/rewriter/src/tests.rs index 7149616b..aeb19b84 100644 --- a/rewriter/src/tests.rs +++ b/rewriter/src/tests.rs @@ -44,6 +44,7 @@ fn assert_steps_from_string_with_domain(domains : Vec<(&str, Domain)>, s: &str) assert!(steps.is_some()); } +#[allow(unused)] fn assert_steps_from_string(s: &str) { assert_steps_from_string_with_domain(vec![], s); } @@ -254,33 +255,34 @@ fn test_log_div_rev_obj() { ]); } -// More +// Folding atoms. #[test] fn test_geo_mean() { - assert_steps_with_domain( + assert_steps_from_string_with_domain( vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], - "(sqrt (mul (var x) (var y)))", - vec![ - ]); + "(neg (sqrt (mul (var x) (var y))))"); } #[test] fn test_quad_over_lin() { - assert_steps_with_domain( + assert_steps_from_string_with_domain( vec![("x", domain::free_dom()), ("y", domain::pos_dom())], - "(div (pow (var x) 2) (var y))", - vec![ - ]); + "(div (pow (var x) 2) (var y))"); } #[test] fn test_norm2() { - assert_steps_with_domain( + assert_steps_from_string_with_domain( vec![("x", domain::free_dom()), ("y", domain::free_dom())], - "(sqrt (add (pow (var x) 2) (pow (var y) 2)))", - vec![ - ]); + "(sqrt (add (pow (var x) 2) (pow (var y) 2)))"); +} + +#[test] +fn test_norm2_with_one() { + assert_steps_from_string_with_domain( + vec![("x", domain::free_dom())], + "(sqrt (add (pow (var x) 2) 1))"); } #[test] @@ -290,18 +292,21 @@ fn test_3_32() { "(div 1 (mul (var x) (var y)))"); } +#[test] +fn test_sqrt_pow4() { + assert_steps_from_string_with_domain( + vec![("x", domain::nonneg_dom())], + "(sqrt (pow (var x) 4))"); +} + #[test] fn test_3_33() { - // NOTE(RFM): The interesting thing here is that we need to somehow have - // 4 -> 2 * 2. Moreover, getting the simplest nodes when matching for - // norm is not enough! - assert_steps_with_domain( - vec![("x", domain::free_dom()), ("y", domain::pos_dom())], - // "(sqrt (add 1 (div (pow (var x) 4) (var y))))", - // "(sqrt (add 1 (div (pow (var x) (mul 2 2)) (var y))))", - "(div (pow (var x) 2) (sqrt (var y)) )", - vec![ - ]); + // NOTE(RFM): x cannot be in the free domain as otherwise we cannot convert + // x^4 to (x^2)^2. + assert_steps_from_string_with_domain( + vec![("x", domain::nonneg_dom()), ("y", domain::pos_dom())], + "(sqrt (add 1 (div (pow (var x) 4) (var y))))" + ); } } From d4f1044a619278262e6461d170e375c8461b538f Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 12:35:27 -0500 Subject: [PATCH 039/189] feat: add rewrite maps --- CvxLean/Tactic/PreDCP/RewriteMapLibrary.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/CvxLean/Tactic/PreDCP/RewriteMapLibrary.lean b/CvxLean/Tactic/PreDCP/RewriteMapLibrary.lean index 42447abd..9bf3ccb1 100644 --- a/CvxLean/Tactic/PreDCP/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/PreDCP/RewriteMapLibrary.lean @@ -181,6 +181,12 @@ register_rewrite_map "div_self" ; "(div ?a ?a)" => "1" := /- Power and square root rules. -/ +register_rewrite_map "one_pow"; "(pow 1 ?a)" => "1" := + simp_or_rw [Real.one_rpow]; + +register_rewrite_map "pow_one"; "(pow ?a 1)" => "?a" := + simp_or_rw [Real.rpow_one]; + register_rewrite_map "pow_add"; "(pow ?a (add ?b ?c))" => "(mul (pow ?a ?b) (pow ?a ?c))" := simp_or_rw [Real.rpow_add (by positivity_ext)]; @@ -223,9 +229,15 @@ register_rewrite_map "one_div_eq_pow_neg_one"; "(div 1 ?a)" => "(pow ?a (neg 1)) register_rewrite_map "sqrt_eq_rpow" ; "(sqrt ?a)" => "(pow ?a 0.5)" := simp_or_rw [Real.sqrt_eq_rpow]; +register_rewrite_map "sqrt_eq_rpow-rev" ; "(pow ?a 0.5)" => "(sqrt ?a)" := + simp_or_rw [←Real.sqrt_eq_rpow]; + register_rewrite_map "pow_half_two"; "(pow (pow ?a 0.5) 2)" => "?a" := simp_or_rw [Real.pow_half_two (by positivity_ext)]; +register_rewrite_map "pow_half_two-rev"; "?a" => "(pow (pow ?a 0.5) 2)" := + simp_or_rw [←Real.pow_half_two (by positivity_ext)]; + /- Exponential and logarithm rules. -/ From d548631c7d5c5a6212aaf17daf2e3261cc5ee33c Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 12:38:39 -0500 Subject: [PATCH 040/189] feat: atom folding rewrite maps --- CvxLean/Tactic/PreDCP/RewriteMapLibrary.lean | 24 +++++++++++++++++--- 1 file changed, 21 insertions(+), 3 deletions(-) diff --git a/CvxLean/Tactic/PreDCP/RewriteMapLibrary.lean b/CvxLean/Tactic/PreDCP/RewriteMapLibrary.lean index 9bf3ccb1..b359fe88 100644 --- a/CvxLean/Tactic/PreDCP/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/PreDCP/RewriteMapLibrary.lean @@ -4,7 +4,7 @@ import CvxLean.Tactic.Util.PositivityExt -- TODO: Move. lemma Real.log_eq_log {x y : ℝ} (hx : 0 < x) (hy : 0 < y) : Real.log x = Real.log y ↔ x = y := - ⟨fun h => by { + ⟨fun h => by have hxmem := Set.mem_Ioi.2 hx have hymem := Set.mem_Ioi.2 hy have heq : Set.restrict (Set.Ioi 0) log ⟨x, hxmem⟩ = @@ -12,8 +12,8 @@ lemma Real.log_eq_log {x y : ℝ} (hx : 0 < x) (hy : 0 < y) : Real.log x = Real. simp [h] have h := Real.log_injOn_pos.injective heq simp [Subtype.eq] at h - exact h - }, fun h => by rw [h]⟩ + exact h, + fun h => by rw [h]⟩ -- TODO: Move. lemma Real.div_pow_eq_mul_pow_neg {a b c : ℝ} (hb : 0 ≤ b) : a / (b ^ c) = a * b ^ (-c) := by @@ -283,4 +283,22 @@ register_rewrite_map "exp_log" ; "(exp (log ?a))" => "?a" := register_rewrite_map "log_exp" ; "(log (exp ?a))" => "?a" := simp_or_rw [Real.log_exp]; + +/- Atom folding. -/ + +register_rewrite_map "xexp_folding"; "(mul ?a (exp ?a))" => "(xexp ?a)" := + rfl; + +register_rewrite_map "entr_folding"; "(neg (mul ?a (log ?a)))" => "(entr ?a)" := + rfl; + +register_rewrite_map "qol_folding"; "(div (pow ?a 2) ?b)" => "(qol ?a ?b)" := + rfl; + +register_rewrite_map "geo_folding"; "(sqrt (mul ?a ?b))" => "(geo ?a ?b)" := + rfl; + +register_rewrite_map "norm2_folding"; "(sqrt (add (pow ?a 2) (pow ?b 2)))" => "(norm2 ?a ?b)" := + rfl; + end CvxLean From 3c0ece8f1ac09f72cc26fb8383d6134ff30eb14a Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 12:39:34 -0500 Subject: [PATCH 041/189] chore: remove `CvxLeanPreDCPTest` --- CvxLeanPreDCPTest.lean | 1 - lakefile.lean | 3 --- 2 files changed, 4 deletions(-) delete mode 100644 CvxLeanPreDCPTest.lean diff --git a/CvxLeanPreDCPTest.lean b/CvxLeanPreDCPTest.lean deleted file mode 100644 index cda5bd24..00000000 --- a/CvxLeanPreDCPTest.lean +++ /dev/null @@ -1 +0,0 @@ -import CvxLean.Test.PreDCP.All diff --git a/lakefile.lean b/lakefile.lean index 5b9c7252..b3c9999a 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -19,9 +19,6 @@ package CvxLean @[default_target] lean_lib CvxLeanTest -@[default_target] -lean_lib CvxLeanPreDCPTest - @[default_target] lean_lib CvxLean From bf1f4d6c52cb45b320d68259aba3ea57244c5461 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 12:48:49 -0500 Subject: [PATCH 042/189] refactor: PreDCP -> Convexify --- CvxLean/Examples/OptimalVehicleSpeed.lean | 32 +++---- .../ChangeOfVariables.lean | 86 +++++++++---------- .../Tactic/{PreDCP => Convexify}/Basic.lean | 0 .../{PreDCP => Convexify}/Convexify.lean | 8 +- CvxLean/Tactic/Convexify/Egg/All.lean | 4 + .../{PreDCP => Convexify}/Egg/EggTypes.lean | 0 .../Egg/FromMinimization.lean | 4 +- .../{PreDCP => Convexify}/Egg/Runner.lean | 2 +- .../{PreDCP => Convexify}/Egg/Sexp.lean | 0 .../{PreDCP => Convexify}/Egg/ToExpr.lean | 2 +- .../{PreDCP => Convexify}/RewriteMapCmd.lean | 18 ++-- .../{PreDCP => Convexify}/RewriteMapExt.lean | 0 .../RewriteMapLibrary.lean | 4 +- .../{PreDCP => Convexify}/UncheckedDCP.lean | 0 CvxLean/Tactic/PreDCP/Egg/All.lean | 4 - CvxLean/Test/All.lean | 3 +- CvxLean/Test/Convexify/All.lean | 5 ++ .../Test/{PreDCP => Convexify}/AlmostDGP.lean | 16 ++-- .../ChangeOfVariables.lean | 36 ++++---- CvxLean/Test/{PreDCP => Convexify}/DGP.lean | 2 +- CvxLean/Test/{PreDCP => Convexify}/DQCP.lean | 4 +- .../{PreDCP => Convexify}/MainExample.lean | 2 +- CvxLean/Test/{PreDCP => Convexify}/Unit.lean | 2 +- .../Test/{PreDCP => Convexify}/almostdgp.py | 0 CvxLean/Test/{PreDCP => Convexify}/dgp.py | 0 CvxLean/Test/{PreDCP => Convexify}/dqcp.py | 0 .../Test/{PreDCP => Convexify}/mainexample.py | 0 CvxLean/Test/PreDCP/All.lean | 5 -- 28 files changed, 120 insertions(+), 119 deletions(-) rename CvxLean/Tactic/{PreDCP => ChangeOfVariables}/ChangeOfVariables.lean (84%) rename CvxLean/Tactic/{PreDCP => Convexify}/Basic.lean (100%) rename CvxLean/Tactic/{PreDCP => Convexify}/Convexify.lean (98%) create mode 100644 CvxLean/Tactic/Convexify/Egg/All.lean rename CvxLean/Tactic/{PreDCP => Convexify}/Egg/EggTypes.lean (100%) rename CvxLean/Tactic/{PreDCP => Convexify}/Egg/FromMinimization.lean (98%) rename CvxLean/Tactic/{PreDCP => Convexify}/Egg/Runner.lean (99%) rename CvxLean/Tactic/{PreDCP => Convexify}/Egg/Sexp.lean (100%) rename CvxLean/Tactic/{PreDCP => Convexify}/Egg/ToExpr.lean (99%) rename CvxLean/Tactic/{PreDCP => Convexify}/RewriteMapCmd.lean (87%) rename CvxLean/Tactic/{PreDCP => Convexify}/RewriteMapExt.lean (100%) rename CvxLean/Tactic/{PreDCP => Convexify}/RewriteMapLibrary.lean (99%) rename CvxLean/Tactic/{PreDCP => Convexify}/UncheckedDCP.lean (100%) delete mode 100644 CvxLean/Tactic/PreDCP/Egg/All.lean create mode 100644 CvxLean/Test/Convexify/All.lean rename CvxLean/Test/{PreDCP => Convexify}/AlmostDGP.lean (92%) rename CvxLean/Test/{PreDCP => Convexify}/ChangeOfVariables.lean (60%) rename CvxLean/Test/{PreDCP => Convexify}/DGP.lean (99%) rename CvxLean/Test/{PreDCP => Convexify}/DQCP.lean (97%) rename CvxLean/Test/{PreDCP => Convexify}/MainExample.lean (95%) rename CvxLean/Test/{PreDCP => Convexify}/Unit.lean (99%) rename CvxLean/Test/{PreDCP => Convexify}/almostdgp.py (100%) rename CvxLean/Test/{PreDCP => Convexify}/dgp.py (100%) rename CvxLean/Test/{PreDCP => Convexify}/dqcp.py (100%) rename CvxLean/Test/{PreDCP => Convexify}/mainexample.py (100%) delete mode 100644 CvxLean/Test/PreDCP/All.lean diff --git a/CvxLean/Examples/OptimalVehicleSpeed.lean b/CvxLean/Examples/OptimalVehicleSpeed.lean index 02a089bd..c0d94895 100644 --- a/CvxLean/Examples/OptimalVehicleSpeed.lean +++ b/CvxLean/Examples/OptimalVehicleSpeed.lean @@ -1,7 +1,7 @@ import CvxLean.Command.Equivalence import CvxLean.Command.Solve import CvxLean.Tactic.Basic.Rename -import CvxLean.Tactic.PreDCP.ChangeOfVariables +import CvxLean.Tactic.ChangeOfVariables.ChangeOfVariables noncomputable section OptimalVehicleSpeed @@ -35,11 +35,11 @@ open FinsetInterval instance [i : Fact (0 < n)] : OfNat (Fin n) 0 := ⟨⟨0, i.out⟩⟩ -def optimalVehicleSpeed (_ : Fact (0 < n)) := +def optimalVehicleSpeed (_ : Fact (0 < n)) := optimization (s : Fin n → ℝ) - minimize ∑ i, (d i / s i) * Φ (s i) - subject to - hsmin : ∀ i, smin i ≤ s i + minimize ∑ i, (d i / s i) * Φ (s i) + subject to + hsmin : ∀ i, smin i ≤ s i hsmax : ∀ i, s i ≤ smax i hτmin : ∀ i, τmin i ≤ ∑ j in [[0, i]], d j / s j hτmax : ∀ i, ∑ j in [[0, i]], d j / s j ≤ τmax i @@ -48,46 +48,46 @@ private lemma simp_fraction : ∀ s : Fin n → ℝ, ∀ i, d i / (d i / s i) = intros s i have h : d i ≠ 0 := by { have d_pos_i := d_pos i - linarith + linarith } rw [← div_mul, div_self h, one_mul] } -equivalence equiv/optimalVehicleSpeedConvex - {n : ℕ} +equivalence equiv/optimalVehicleSpeedConvex + {n : ℕ} (n_pos : 0 < n) (d : Fin n → ℝ) (d_pos : ∀ i, 0 < d i) (τmin τmax : Fin n → ℝ) (smin smax : Fin n → ℝ) (smin_pos : ∀ i, 0 < smin i) - (Φ : ℝ → ℝ) : - optimalVehicleSpeed d τmin τmax smin smax Φ ⟨n_pos⟩ := by + (Φ : ℝ → ℝ) : + optimalVehicleSpeed d τmin τmax smin smax Φ ⟨n_pos⟩ := by unfold optimalVehicleSpeed apply Minimization.Equivalence.trans apply ChangeOfVariables.toEquivalence (fun t => d / t) { rintro s ⟨hsmin, _hsmax, _hτmin, _hτmax⟩ i split_ands { --suffices h : 0 < s i by exact (ne_of_lt h).symm - --exact lt_of_lt_of_le (smin_pos i) (hsmin i) + --exact lt_of_lt_of_le (smin_pos i) (hsmin i) have smin_pos_i := smin_pos i have hsmin_i := hsmin i linarith } - { --exact (ne_of_lt (d_pos i)).symm + { --exact (ne_of_lt (d_pos i)).symm have d_pos_i := d_pos i linarith } } - -- IDEA: This can be done by change of variables by detecting that the + -- IDEA: This can be done by change of variables by detecting that the -- variable is a vector. apply Minimization.Equivalence.trans apply MinimizationQ.rewrite_objective { rintro s ⟨_hsmin, _hsmax, _hτmin, _hτmax⟩ simp only [Pi.div_apply, simp_fraction d d_pos s] - rfl + rfl } apply Minimization.Equivalence.trans - apply MinimizationQ.rewrite_constraint_3 + apply MinimizationQ.rewrite_constraint_3 { rintro s _hsmin _hsmax _hτmax simp only [Pi.div_apply, simp_fraction d d_pos s] rfl @@ -108,7 +108,7 @@ equivalence equiv/optimalVehicleSpeedConvex -- TODO: rename_vars in equivalence mode. -- IDEA: simplify_objective [...], simplify_constraint i [...], --- Also rewrite_objective and rewrite_constraint. Potentially get rid of the +-- Also rewrite_objective and rewrite_constraint. Potentially get rid of the -- rewrite_constraint_x lemmas and handle that in meta. end OptimalVehicleSpeed diff --git a/CvxLean/Tactic/PreDCP/ChangeOfVariables.lean b/CvxLean/Tactic/ChangeOfVariables/ChangeOfVariables.lean similarity index 84% rename from CvxLean/Tactic/PreDCP/ChangeOfVariables.lean rename to CvxLean/Tactic/ChangeOfVariables/ChangeOfVariables.lean index 813e754d..c5768db2 100644 --- a/CvxLean/Tactic/PreDCP/ChangeOfVariables.lean +++ b/CvxLean/Tactic/ChangeOfVariables/ChangeOfVariables.lean @@ -3,25 +3,25 @@ import CvxLean.Lib.Minimization import CvxLean.Lib.Equivalence import CvxLean.Meta.Minimization import CvxLean.Tactic.Conv.ConvOpt -import CvxLean.Tactic.DCP.AtomLibrary -import CvxLean.Tactic.PreDCP.Convexify -- TODO(RFM): we only need normNumCleanUp from here, factor out. +import CvxLean.Tactic.DCP.AtomLibrary.All +import CvxLean.Tactic.Convexify.Convexify -- TODO(RFM): we only need normNumCleanUp from here, factor out. namespace CvxLean open Minimization -class ChangeOfVariables {D E} (c : E → D) where +class ChangeOfVariables {D E} (c : E → D) where inv : D → E - condition : D → Prop + condition : D → Prop property : ∀ x, condition x → c (inv x) = x -def ChangeOfVariables.toEquivalence {D E R} [Preorder R] +def ChangeOfVariables.toEquivalence {D E R} [Preorder R] {f : D → R} {cs : D → Prop} (c : E → D) [cov : ChangeOfVariables c] (h : ∀ x, cs x → cov.condition x) : - Equivalence + Equivalence (Minimization.mk f cs) - (Minimization.mk (fun x => f (c x)) (fun x => cs (c x))) := + (Minimization.mk (fun x => f (c x)) (fun x => cs (c x))) := StrongEquivalence.toEquivalence <| { phi := fun x => cov.inv x psi := fun y => c y @@ -33,8 +33,8 @@ def ChangeOfVariables.toEquivalence {D E R} [Preorder R] section Structural -instance ChangeOfVariables.comp {D E F} (c₁ : E → D) (c₂ : F → E) - [cov₁ : ChangeOfVariables c₁] [cov₂ : ChangeOfVariables c₂] : +instance ChangeOfVariables.comp {D E F} (c₁ : E → D) (c₂ : F → E) + [cov₁ : ChangeOfVariables c₁] [cov₂ : ChangeOfVariables c₂] : ChangeOfVariables (c₁ ∘ c₂) := { inv := cov₂.inv ∘ cov₁.inv condition := fun x => cov₁.condition x ∧ cov₂.condition (cov₁.inv x) @@ -42,19 +42,19 @@ instance ChangeOfVariables.comp {D E F} (c₁ : E → D) (c₂ : F → E) simp [cov₂.property (cov₁.inv x) hx₂] simp [cov₁.property x hx₁] } - } + } -instance ChangeOfVariables.prod_left {D E F} (c : E → D) - [cov : ChangeOfVariables c] : +instance ChangeOfVariables.prod_left {D E F} (c : E → D) + [cov : ChangeOfVariables c] : ChangeOfVariables (fun x : E × F => (c x.1, x.2)) := { inv := fun ⟨x₁, x₂⟩ => (cov.inv x₁, x₂) condition := fun ⟨x₁, _⟩ => cov.condition x₁ property := fun ⟨x₁, x₂⟩ hx => by simp [cov.property x₁ hx] } -instance ChangeOfVariables.prod_left_binary_left_constant +instance ChangeOfVariables.prod_left_binary_left_constant {D E F} {T} (c : T → E → D) (t : T) - [cov : ChangeOfVariables (fun x => c t x)] : + [cov : ChangeOfVariables (fun x => c t x)] : ChangeOfVariables (fun x : E × F => (c t x.1, x.2)) := { inv := fun ⟨x₁, x₂⟩ => (cov.inv x₁, x₂) condition := fun ⟨x₁, _⟩ => cov.condition x₁ @@ -63,7 +63,7 @@ instance ChangeOfVariables.prod_left_binary_left_constant instance ChangeOfVariables.prod_left_binary_right_constant {D E F} {T} (c : E → T → D) (t : T) - [cov : ChangeOfVariables (fun x => c x t)] : + [cov : ChangeOfVariables (fun x => c x t)] : ChangeOfVariables (fun x : E × F => (c x.1 t, x.2)) := { inv := fun ⟨x₁, x₂⟩ => (cov.inv x₁, x₂) condition := fun ⟨x₁, _⟩ => cov.condition x₁ @@ -71,7 +71,7 @@ instance ChangeOfVariables.prod_left_binary_right_constant } instance ChangeOfVariables.prod_right {D E F} (c : E → D) - [cov : ChangeOfVariables c] : + [cov : ChangeOfVariables c] : ChangeOfVariables (fun x : F × E => (x.1, c x.2)) := { inv := fun ⟨x₁, x₂⟩ => (x₁, cov.inv x₂) condition := fun ⟨_, x₂⟩ => cov.condition x₂ @@ -80,7 +80,7 @@ instance ChangeOfVariables.prod_right {D E F} (c : E → D) instance ChangeOfVariables.prod_right_binary_left_constant {D E F} {T} (c : T → E → D) (t : T) - [cov : ChangeOfVariables (fun x => c t x)] : + [cov : ChangeOfVariables (fun x => c t x)] : ChangeOfVariables (fun x : F × E => (x.1, c t x.2)) := { inv := fun ⟨x₁, x₂⟩ => (x₁, cov.inv x₂) condition := fun ⟨_, x₂⟩ => cov.condition x₂ @@ -89,7 +89,7 @@ instance ChangeOfVariables.prod_right_binary_left_constant instance ChangeOfVariables.prod_right_binary_right_constant {D E F} {T} (c : E → T → D) (t : T) - [cov : ChangeOfVariables (fun x => c x t)] : + [cov : ChangeOfVariables (fun x => c x t)] : ChangeOfVariables (fun x : F × E => (x.1, c x.2 t)) := { inv := fun ⟨x₁, x₂⟩ => (x₁, cov.inv x₂) condition := fun ⟨_, x₂⟩ => cov.condition x₂ @@ -110,17 +110,17 @@ instance : ChangeOfVariables Real.exp := { inv := Real.log condition := fun x => 0 < x property := fun _ hx => Real.exp_log hx -} +} --- NOTE(RFM): x ≠ 0 is technically not necessary as division is defined on all --- of ℝ, but we want to avoid division by zero. +-- NOTE(RFM): x ≠ 0 is technically not necessary as division is defined on all +-- of ℝ, but we want to avoid division by zero. instance : ChangeOfVariables (fun x : ℝ => x⁻¹) := { inv := fun x => x⁻¹ - condition := fun x => x ≠ 0 + condition := fun x => x ≠ 0 property := fun x _ => by field_simp } --- NOTE(RFM): a ≠ 0 is not given as a parameter but instead added to the +-- NOTE(RFM): a ≠ 0 is not given as a parameter but instead added to the -- condition to make type class inference work. instance {a : ℝ} : ChangeOfVariables (fun x : ℝ => a * x) := { inv := fun x => (1 / a) * x @@ -144,7 +144,7 @@ end RealInstances noncomputable section VecInstances -instance {n : ℕ} {a : Fin n → ℝ} : +instance {n : ℕ} {a : Fin n → ℝ} : ChangeOfVariables (fun (v : Fin n → ℝ) => a / v) := { inv := fun v => a / v condition := fun v => ∀ i, v i ≠ 0 ∧ a i ≠ 0 @@ -161,9 +161,9 @@ The idea here is to have a tactic change_of_variables (u) (x ↦ e^u) -1. Detect exactly where to apply the change of variables. +1. Detect exactly where to apply the change of variables. 2. Syntesize the instance. -2. Prove the conditions. +2. Prove the conditions. 3. Apply the c-of-v to equivalence theorem. For now, it only works with real variables. @@ -173,20 +173,20 @@ section Tactic open Lean Elab Meta Tactic Term -syntax (name := change_of_variables) +syntax (name := change_of_variables) "change_of_variables" "(" ident ")" "(" ident "↦" term ")" : tactic @[tactic change_of_variables] def evalChangeOfVariables : Tactic := fun stx => match stx with - | `(tactic| change_of_variables ($newVarStx) ($varToChangeStx ↦ $changeStx)) => + | `(tactic| change_of_variables ($newVarStx) ($varToChangeStx ↦ $changeStx)) => withMainContext do let newVar := newVarStx.getId let varToChange := varToChangeStx.getId let gTarget ← getMainTarget - if !gTarget.isAppOf ``Minimization.Equivalence then + if !gTarget.isAppOf ``Minimization.Equivalence then throwError "`change_of_variables` can only be applied to equivalences." - + -- TODO(RFM): Only working with equivalence for now. let gExprRaw ← liftM <| Meta.getExprRawFromGoal true gTarget let gExpr ← MinimizationExpr.fromExpr gExprRaw @@ -194,12 +194,12 @@ def evalChangeOfVariables : Tactic := fun stx => match stx with -- Find change of variables location. let covIdx := vars.findIdx? (fun ⟨n, _⟩ => n == varToChange) - if covIdx.isNone then + if covIdx.isNone then throwError "Variable {varToChange} not found in domain." let covIdx := covIdx.get! -- New domain. - let newVars := vars.map (fun ⟨n, ty⟩ => ⟨if n = varToChange then newVar else n, ty⟩) + let newVars := vars.map (fun ⟨n, ty⟩ => ⟨if n = varToChange then newVar else n, ty⟩) let newDomain := composeDomain newVars -- Construct change of variables function. @@ -208,23 +208,23 @@ def evalChangeOfVariables : Tactic := fun stx => match stx with let changeFnStx ← `(fun $newVarStx => $changeStx) let changeFn ← Tactic.elabTerm changeFnStx none -- c(x) - let changeTerm ← Core.betaReduce <| + let changeTerm ← Core.betaReduce <| mkApp changeFn (mkFVar (FVarId.mk varToChange)) -- (x₁, ..., u, ..., xₙ) ↦ (x₁, ..., c(u), ..., xₙ) let c ← withLocalDeclD `p newDomain fun p => do Meta.withDomainLocalDecls newDomain p fun xs prs => do -- (x₁, ..., c(xᵢ), ..., xₙ) let fullChangeTerm ← Expr.mkProd <| - (xs.take covIdx) ++ - #[changeTerm] ++ + (xs.take covIdx) ++ + #[changeTerm] ++ (xs.drop (covIdx + 1)) let replacedFVars := Expr.replaceFVars fullChangeTerm fvars xs mkLambdaFVars #[p] (Expr.replaceFVars replacedFVars xs prs) - + -- Apply transitivity. let g ← getMainGoal let gsTrans ← evalTacticAt (← `(tactic| apply Minimization.Equivalence.trans)) g - if gsTrans.length != 4 then + if gsTrans.length != 4 then throwError "Equivalence mode. Expected 4 goals after applying `trans`, got {gsTrans.length}." let gToChange := gsTrans[0]! let gNext := gsTrans[1]! @@ -237,24 +237,24 @@ def evalChangeOfVariables : Tactic := fun stx => match stx with (mkAppN (mkConst ``Preorder [levelZero]) #[R]) let f := gExpr.objFun let cs := gExpr.constraints - -- let cov ← mkFreshExprMVar <| + -- let cov ← mkFreshExprMVar <| -- (mkAppN (mkConst ``ChangeOfVariables [levelZero, levelZero]) #[D, E, c]) let toApply := mkAppN (mkConst ``ChangeOfVariables.toEquivalence) #[D, E, R, RPreorder, f, cs, c] let gsAfterApply ← gToChange.apply toApply - if gsAfterApply.length != 1 then + if gsAfterApply.length != 1 then throwError ( - "Failed to apply `ChangeOfVariables.toEquivalence`. " ++ + "Failed to apply `ChangeOfVariables.toEquivalence`. " ++ "Make sure that the change of variables is inferrable by type class resolution.") -- Solve change of variables condition. let gCondition := gsAfterApply[0]! let (_, gCondition) ← gCondition.intros - let gsFinal ← evalTacticAt + let gsFinal ← evalTacticAt (← `(tactic| simp [ChangeOfVariables.condition] <;> intros <;> positivity)) gCondition - if gsFinal.length != 0 then - for g in gsFinal do + if gsFinal.length != 0 then + for g in gsFinal do dbg_trace s!"Could not prove {← Meta.ppGoal g}." throwError "Failed to solve change of variables condition." diff --git a/CvxLean/Tactic/PreDCP/Basic.lean b/CvxLean/Tactic/Convexify/Basic.lean similarity index 100% rename from CvxLean/Tactic/PreDCP/Basic.lean rename to CvxLean/Tactic/Convexify/Basic.lean diff --git a/CvxLean/Tactic/PreDCP/Convexify.lean b/CvxLean/Tactic/Convexify/Convexify.lean similarity index 98% rename from CvxLean/Tactic/PreDCP/Convexify.lean rename to CvxLean/Tactic/Convexify/Convexify.lean index 216059b2..9b76c851 100644 --- a/CvxLean/Tactic/PreDCP/Convexify.lean +++ b/CvxLean/Tactic/Convexify/Convexify.lean @@ -2,10 +2,10 @@ import Lean import Mathlib.Tactic.NormNum import CvxLean.Meta.Minimization import CvxLean.Lib.Equivalence -import CvxLean.Tactic.PreDCP.Basic -import CvxLean.Tactic.PreDCP.RewriteMapExt -import CvxLean.Tactic.PreDCP.RewriteMapLibrary -import CvxLean.Tactic.PreDCP.Egg.All +import CvxLean.Tactic.Convexify.Basic +import CvxLean.Tactic.Convexify.RewriteMapExt +import CvxLean.Tactic.Convexify.RewriteMapLibrary +import CvxLean.Tactic.Convexify.Egg.All namespace CvxLean diff --git a/CvxLean/Tactic/Convexify/Egg/All.lean b/CvxLean/Tactic/Convexify/Egg/All.lean new file mode 100644 index 00000000..fd58f321 --- /dev/null +++ b/CvxLean/Tactic/Convexify/Egg/All.lean @@ -0,0 +1,4 @@ +import CvxLean.Tactic.Convexify.Egg.Sexp +import CvxLean.Tactic.Convexify.Egg.ToExpr +import CvxLean.Tactic.Convexify.Egg.FromMinimization +import CvxLean.Tactic.Convexify.Egg.Runner diff --git a/CvxLean/Tactic/PreDCP/Egg/EggTypes.lean b/CvxLean/Tactic/Convexify/Egg/EggTypes.lean similarity index 100% rename from CvxLean/Tactic/PreDCP/Egg/EggTypes.lean rename to CvxLean/Tactic/Convexify/Egg/EggTypes.lean diff --git a/CvxLean/Tactic/PreDCP/Egg/FromMinimization.lean b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean similarity index 98% rename from CvxLean/Tactic/PreDCP/Egg/FromMinimization.lean rename to CvxLean/Tactic/Convexify/Egg/FromMinimization.lean index 36410c11..efc04635 100644 --- a/CvxLean/Tactic/PreDCP/Egg/FromMinimization.lean +++ b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean @@ -1,6 +1,6 @@ import CvxLean.Tactic.DCP.Tree -import CvxLean.Tactic.PreDCP.UncheckedDCP -import CvxLean.Tactic.PreDCP.Egg.EggTypes +import CvxLean.Tactic.Convexify.UncheckedDCP +import CvxLean.Tactic.Convexify.Egg.EggTypes namespace CvxLean diff --git a/CvxLean/Tactic/PreDCP/Egg/Runner.lean b/CvxLean/Tactic/Convexify/Egg/Runner.lean similarity index 99% rename from CvxLean/Tactic/PreDCP/Egg/Runner.lean rename to CvxLean/Tactic/Convexify/Egg/Runner.lean index 59ec9344..8e1fe199 100644 --- a/CvxLean/Tactic/PreDCP/Egg/Runner.lean +++ b/CvxLean/Tactic/Convexify/Egg/Runner.lean @@ -1,5 +1,5 @@ import Lean -import CvxLean.Tactic.PreDCP.Egg.EggTypes +import CvxLean.Tactic.Convexify.Egg.EggTypes open Lean diff --git a/CvxLean/Tactic/PreDCP/Egg/Sexp.lean b/CvxLean/Tactic/Convexify/Egg/Sexp.lean similarity index 100% rename from CvxLean/Tactic/PreDCP/Egg/Sexp.lean rename to CvxLean/Tactic/Convexify/Egg/Sexp.lean diff --git a/CvxLean/Tactic/PreDCP/Egg/ToExpr.lean b/CvxLean/Tactic/Convexify/Egg/ToExpr.lean similarity index 99% rename from CvxLean/Tactic/PreDCP/Egg/ToExpr.lean rename to CvxLean/Tactic/Convexify/Egg/ToExpr.lean index be32c9a2..b5bca3f1 100644 --- a/CvxLean/Tactic/PreDCP/Egg/ToExpr.lean +++ b/CvxLean/Tactic/Convexify/Egg/ToExpr.lean @@ -1,6 +1,6 @@ import CvxLean.Lib.Math.Data.Real import CvxLean.Meta.Minimization -import CvxLean.Tactic.PreDCP.Egg.Sexp +import CvxLean.Tactic.Convexify.Egg.Sexp import CvxLean.Tactic.DCP.Tree namespace CvxLean diff --git a/CvxLean/Tactic/PreDCP/RewriteMapCmd.lean b/CvxLean/Tactic/Convexify/RewriteMapCmd.lean similarity index 87% rename from CvxLean/Tactic/PreDCP/RewriteMapCmd.lean rename to CvxLean/Tactic/Convexify/RewriteMapCmd.lean index 9ff29287..559be670 100644 --- a/CvxLean/Tactic/PreDCP/RewriteMapCmd.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapCmd.lean @@ -1,33 +1,33 @@ import Lean import Std.Linter.UnreachableTactic -import CvxLean.Tactic.PreDCP.RewriteMapExt +import CvxLean.Tactic.Convexify.RewriteMapExt namespace CvxLean open Lean Parser -syntax (name := registerRewriteMap) +syntax (name := registerRewriteMap) "register_rewrite_map " str " ; " str " => " str " := " tactic ";" : command -syntax (name := registerRewriteMapBidirectional) +syntax (name := registerRewriteMapBidirectional) "register_rewrite_map " str " ; " str " <=> " str " := " tactic ";" : command -macro_rules - | `(register_rewrite_map $rwName ; $rwTarget <=> $rwGoal := $tac;) => +macro_rules + | `(register_rewrite_map $rwName ; $rwTarget <=> $rwGoal := $tac;) => if let some rwNameStr := Syntax.isStrLit? rwName then let rwNameRev := Syntax.mkStrLit (rwNameStr ++ "-rev") `(register_rewrite_map $rwName ; $rwTarget => $rwGoal := $tac; register_rewrite_map $rwNameRev ; $rwGoal => $rwTarget := $tac;) else `(throwError "register_rewrite_map error: expected string") -syntax (name := registerObjFunRewriteMap) +syntax (name := registerObjFunRewriteMap) "register_objFun_rewrite_map " str " ; " str " => " str " := " tactic ";" : command open Lean.Elab Lean.Elab.Command set_option linter.unreachableTactic false in @[command_elab registerRewriteMap] def elabRegisterEggRewrite : CommandElab -| `(register_rewrite_map $rwName ; $rwTarget => $rwGoal := $tac;) => do +| `(register_rewrite_map $rwName ; $rwTarget => $rwGoal := $tac;) => do let rwNameStr := rwName.getString -- NOTE: We ignore this for now. let _rwTargetStr := rwTarget.getString @@ -37,7 +37,7 @@ set_option linter.unreachableTactic false in set_option linter.unreachableTactic false in @[command_elab registerObjFunRewriteMap] def elabRegisterObjFunEggRewrite : CommandElab -| `(register_objFun_rewrite_map $rwName ; $rwTarget => $rwGoal := $tac;) => do +| `(register_objFun_rewrite_map $rwName ; $rwTarget => $rwGoal := $tac;) => do let rwNameStr := rwName.getString -- NOTE: We ignore this for now. let _rwTargetStr := rwTarget.getString @@ -45,4 +45,4 @@ set_option linter.unreachableTactic false in liftTermElabM <| addRewriteMapEntry rwNameStr tac true | _ => throwUnsupportedSyntax -end CvxLean +end CvxLean diff --git a/CvxLean/Tactic/PreDCP/RewriteMapExt.lean b/CvxLean/Tactic/Convexify/RewriteMapExt.lean similarity index 100% rename from CvxLean/Tactic/PreDCP/RewriteMapExt.lean rename to CvxLean/Tactic/Convexify/RewriteMapExt.lean diff --git a/CvxLean/Tactic/PreDCP/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean similarity index 99% rename from CvxLean/Tactic/PreDCP/RewriteMapLibrary.lean rename to CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index b359fe88..76f0d8ac 100644 --- a/CvxLean/Tactic/PreDCP/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -1,5 +1,5 @@ -import CvxLean.Tactic.PreDCP.RewriteMapCmd -import CvxLean.Tactic.PreDCP.Basic +import CvxLean.Tactic.Convexify.RewriteMapCmd +import CvxLean.Tactic.Convexify.Basic import CvxLean.Tactic.Util.PositivityExt -- TODO: Move. diff --git a/CvxLean/Tactic/PreDCP/UncheckedDCP.lean b/CvxLean/Tactic/Convexify/UncheckedDCP.lean similarity index 100% rename from CvxLean/Tactic/PreDCP/UncheckedDCP.lean rename to CvxLean/Tactic/Convexify/UncheckedDCP.lean diff --git a/CvxLean/Tactic/PreDCP/Egg/All.lean b/CvxLean/Tactic/PreDCP/Egg/All.lean deleted file mode 100644 index ff80ab0c..00000000 --- a/CvxLean/Tactic/PreDCP/Egg/All.lean +++ /dev/null @@ -1,4 +0,0 @@ -import CvxLean.Tactic.PreDCP.Egg.Sexp -import CvxLean.Tactic.PreDCP.Egg.ToExpr -import CvxLean.Tactic.PreDCP.Egg.FromMinimization -import CvxLean.Tactic.PreDCP.Egg.Runner diff --git a/CvxLean/Test/All.lean b/CvxLean/Test/All.lean index 942e8a61..213e03c3 100644 --- a/CvxLean/Test/All.lean +++ b/CvxLean/Test/All.lean @@ -1,3 +1,4 @@ +import CvxLean.Test.Convexify.All import CvxLean.Test.DCP.All import CvxLean.Test.Problems.All -import CvxLean.Test.Other.All +import CvxLean.Test.Other.All diff --git a/CvxLean/Test/Convexify/All.lean b/CvxLean/Test/Convexify/All.lean new file mode 100644 index 00000000..0fd3f819 --- /dev/null +++ b/CvxLean/Test/Convexify/All.lean @@ -0,0 +1,5 @@ +import CvxLean.Test.Convexify.MainExample +import CvxLean.Test.Convexify.Unit +import CvxLean.Test.Convexify.DGP +import CvxLean.Test.Convexify.AlmostDGP +import CvxLean.Test.Convexify.DQCP diff --git a/CvxLean/Test/PreDCP/AlmostDGP.lean b/CvxLean/Test/Convexify/AlmostDGP.lean similarity index 92% rename from CvxLean/Test/PreDCP/AlmostDGP.lean rename to CvxLean/Test/Convexify/AlmostDGP.lean index 03bc0998..1a9fc408 100644 --- a/CvxLean/Test/PreDCP/AlmostDGP.lean +++ b/CvxLean/Test/Convexify/AlmostDGP.lean @@ -1,5 +1,5 @@ import CvxLean.Command.Solve -import CvxLean.Tactic.PreDCP.Convexify +import CvxLean.Tactic.Convexify.Convexify import CvxLean.Test.Util.TimeCmd namespace AlmostGP @@ -24,7 +24,7 @@ time_cmd reduction reda1/dcpa1 : agp1 := by #print dcpa1 -- def dcpa1 : Minimization ℝ ℝ := --- optimization (x : ℝ) +-- optimization (x : ℝ) -- minimize x -- subject to -- h2 : exp (x * 2) - 10123 / 1000 ≤ 0 @@ -42,7 +42,7 @@ def agp2 := subject to h1 : 0 < x h2 : 0 < y - h3 : x * y - 5.382 ≤ 0 + h3 : x * y - 5.382 ≤ 0 time_cmd reduction reda2/dcpa2 : agp2 := by map_exp @@ -50,7 +50,7 @@ time_cmd reduction reda2/dcpa2 : agp2 := by #print dcpa2 -- def dcpa2 : Minimization (ℝ × ℝ) ℝ := --- optimization (x : ℝ) (y : ℝ) +-- optimization (x : ℝ) (y : ℝ) -- minimize x -- subject to -- h3 : exp (x + y) - 2691 / 500 ≤ 0 @@ -80,14 +80,14 @@ time_cmd reduction reda3/dcpa3 : agp3 := by #print dcpa3 -- def dcpa3 : Minimization (ℝ × ℝ × ℝ) ℝ := --- optimization (x : ℝ) (y : ℝ) (z : ℝ) +-- optimization (x : ℝ) (y : ℝ) (z : ℝ) -- minimize exp y + (exp x + exp z) -- subject to -- h4 : log 2 ≤ x -- h5 : x ≤ log 3 -- h6 : 6 * exp (y - z - x * 2) ≤ 1 - exp (x * -(3 / 2)) -- h7 : x + y = z - + solve dcpa3 end AlmostDGP3 @@ -109,14 +109,14 @@ time_cmd reduction reda4/dcpa4 : agp4 := by #print dcpa4 -- def dcpa4 : Minimization (ℝ × ℝ) ℝ := --- optimization (x : ℝ) (y : ℝ) +-- optimization (x : ℝ) (y : ℝ) -- minimize -(x + y) -- subject to -- h3 : exp x ≤ 2 - exp (x + y) - exp y solve dcpa4 -end AlmostDGP4 +end AlmostDGP4 -- /- This problem is not convex. -/ -- section AlmostDGP5 diff --git a/CvxLean/Test/PreDCP/ChangeOfVariables.lean b/CvxLean/Test/Convexify/ChangeOfVariables.lean similarity index 60% rename from CvxLean/Test/PreDCP/ChangeOfVariables.lean rename to CvxLean/Test/Convexify/ChangeOfVariables.lean index 8a9338eb..d92d11b6 100644 --- a/CvxLean/Test/PreDCP/ChangeOfVariables.lean +++ b/CvxLean/Test/Convexify/ChangeOfVariables.lean @@ -1,61 +1,61 @@ import CvxLean.Command.Equivalence -import CvxLean.Tactic.PreDCP.ChangeOfVariables +import CvxLean.Tactic.ChangeOfVariables.ChangeOfVariables open CvxLean noncomputable section ChangeOfVariablesInstances -instance : ChangeOfVariables (fun (x : ℝ) => Real.exp x) := by +instance : ChangeOfVariables (fun (x : ℝ) => Real.exp x) := by infer_instance -instance : ChangeOfVariables (fun (x : ℝ) => x + 1) := by +instance : ChangeOfVariables (fun (x : ℝ) => x + 1) := by infer_instance -instance : ChangeOfVariables (fun (x : ℝ) => x - 1) := by +instance : ChangeOfVariables (fun (x : ℝ) => x - 1) := by infer_instance -instance : ChangeOfVariables (fun (x : ℝ) => x⁻¹) := by +instance : ChangeOfVariables (fun (x : ℝ) => x⁻¹) := by infer_instance -instance : ChangeOfVariables (fun (x : ℝ) => 3 * x) := by +instance : ChangeOfVariables (fun (x : ℝ) => 3 * x) := by infer_instance -instance : ChangeOfVariables (fun (x : ℝ) => 3 / x) := by +instance : ChangeOfVariables (fun (x : ℝ) => 3 / x) := by infer_instance -instance : ChangeOfVariables (fun (xy : ℝ × ℝ) => (xy.1 + 1, xy.2)) := by +instance : ChangeOfVariables (fun (xy : ℝ × ℝ) => (xy.1 + 1, xy.2)) := by infer_instance -instance : ChangeOfVariables (fun (xy : ℝ × ℝ) => (Real.exp xy.1, xy.2)) := by +instance : ChangeOfVariables (fun (xy : ℝ × ℝ) => (Real.exp xy.1, xy.2)) := by infer_instance -instance : ChangeOfVariables (fun (xy : ℝ × ℝ) => (xy.1 + 1, xy.2)) := by +instance : ChangeOfVariables (fun (xy : ℝ × ℝ) => (xy.1 + 1, xy.2)) := by infer_instance end ChangeOfVariablesInstances -def p := +def p := optimization (x y : ℝ) minimize (x + y) - subject to + subject to h : 0 < x -equivalence eqv/q1 : p := by - unfold p +equivalence eqv/q1 : p := by + unfold p change_of_variables (u) (x ↦ Real.exp u) equivalence_rfl #print q1 -equivalence eqv2/q2 : p := by - unfold p +equivalence eqv2/q2 : p := by + unfold p change_of_variables (u) (y ↦ u + (1 : ℝ)) equivalence_rfl #print q2 -equivalence eqv3/q3 : p := by - unfold p +equivalence eqv3/q3 : p := by + unfold p change_of_variables (u) (x ↦ (1 : ℝ) / u) equivalence_rfl diff --git a/CvxLean/Test/PreDCP/DGP.lean b/CvxLean/Test/Convexify/DGP.lean similarity index 99% rename from CvxLean/Test/PreDCP/DGP.lean rename to CvxLean/Test/Convexify/DGP.lean index 8163f516..7c6ef406 100644 --- a/CvxLean/Test/PreDCP/DGP.lean +++ b/CvxLean/Test/Convexify/DGP.lean @@ -1,5 +1,5 @@ import CvxLean.Command.Solve -import CvxLean.Tactic.PreDCP.Convexify +import CvxLean.Tactic.Convexify.Convexify import CvxLean.Test.Util.TimeCmd namespace GP diff --git a/CvxLean/Test/PreDCP/DQCP.lean b/CvxLean/Test/Convexify/DQCP.lean similarity index 97% rename from CvxLean/Test/PreDCP/DQCP.lean rename to CvxLean/Test/Convexify/DQCP.lean index 4224b07f..80ff7e57 100644 --- a/CvxLean/Test/PreDCP/DQCP.lean +++ b/CvxLean/Test/Convexify/DQCP.lean @@ -1,7 +1,7 @@ import CvxLean.Command.Solve import CvxLean.Command.Equivalence -import CvxLean.Tactic.PreDCP.Basic -import CvxLean.Tactic.PreDCP.Convexify +import CvxLean.Tactic.Convexify.Basic +import CvxLean.Tactic.Convexify.Convexify import CvxLean.Test.Util.TimeCmd -- TODO: Move. diff --git a/CvxLean/Test/PreDCP/MainExample.lean b/CvxLean/Test/Convexify/MainExample.lean similarity index 95% rename from CvxLean/Test/PreDCP/MainExample.lean rename to CvxLean/Test/Convexify/MainExample.lean index 00137582..7add17cb 100644 --- a/CvxLean/Test/PreDCP/MainExample.lean +++ b/CvxLean/Test/Convexify/MainExample.lean @@ -1,6 +1,6 @@ import CvxLean.Command.Solve import CvxLean.Command.Equivalence -import CvxLean.Tactic.PreDCP.Convexify +import CvxLean.Tactic.Convexify.Convexify import CvxLean.Test.Util.TimeCmd namespace MainExample diff --git a/CvxLean/Test/PreDCP/Unit.lean b/CvxLean/Test/Convexify/Unit.lean similarity index 99% rename from CvxLean/Test/PreDCP/Unit.lean rename to CvxLean/Test/Convexify/Unit.lean index 9b689117..d32aa8cf 100644 --- a/CvxLean/Test/PreDCP/Unit.lean +++ b/CvxLean/Test/Convexify/Unit.lean @@ -1,6 +1,6 @@ import CvxLean.Command.Solve import CvxLean.Command.Equivalence -import CvxLean.Tactic.PreDCP.Convexify +import CvxLean.Tactic.Convexify.Convexify import CvxLean.Tactic.DCP.Dcp import CvxLean.Test.Util.TimeCmd diff --git a/CvxLean/Test/PreDCP/almostdgp.py b/CvxLean/Test/Convexify/almostdgp.py similarity index 100% rename from CvxLean/Test/PreDCP/almostdgp.py rename to CvxLean/Test/Convexify/almostdgp.py diff --git a/CvxLean/Test/PreDCP/dgp.py b/CvxLean/Test/Convexify/dgp.py similarity index 100% rename from CvxLean/Test/PreDCP/dgp.py rename to CvxLean/Test/Convexify/dgp.py diff --git a/CvxLean/Test/PreDCP/dqcp.py b/CvxLean/Test/Convexify/dqcp.py similarity index 100% rename from CvxLean/Test/PreDCP/dqcp.py rename to CvxLean/Test/Convexify/dqcp.py diff --git a/CvxLean/Test/PreDCP/mainexample.py b/CvxLean/Test/Convexify/mainexample.py similarity index 100% rename from CvxLean/Test/PreDCP/mainexample.py rename to CvxLean/Test/Convexify/mainexample.py diff --git a/CvxLean/Test/PreDCP/All.lean b/CvxLean/Test/PreDCP/All.lean deleted file mode 100644 index e94de308..00000000 --- a/CvxLean/Test/PreDCP/All.lean +++ /dev/null @@ -1,5 +0,0 @@ -import CvxLean.Test.PreDCP.MainExample -import CvxLean.Test.PreDCP.Unit -import CvxLean.Test.PreDCP.DGP -import CvxLean.Test.PreDCP.AlmostDGP -import CvxLean.Test.PreDCP.DQCP From 56fdbba30e0e28f1173f2d85b1cfc991acc5edc4 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 13:00:06 -0500 Subject: [PATCH 043/189] fix: only pull out-of-context term if no context found --- rewriter/src/extract.rs | 8 +++++--- 1 file changed, 5 insertions(+), 3 deletions(-) diff --git a/rewriter/src/extract.rs b/rewriter/src/extract.rs index f4e0a0f2..348441db 100644 --- a/rewriter/src/extract.rs +++ b/rewriter/src/extract.rs @@ -60,9 +60,11 @@ fn get_step_aux( *expected_term = Some(c_s); } _ => { - // Out-of-context extraction. Useful for testing. - let c_s = next.get_recexpr().to_string(); - *expected_term = Some(c_s); + if expected_term.is_none() { + // Out-of-context extraction. Useful for testing. + let c_s = next.get_recexpr().to_string(); + *expected_term = Some(c_s); + } } } From 579c5ff72704ec6522fa7d1d8dae7931c2d91e2d Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 13:00:31 -0500 Subject: [PATCH 044/189] fix: rules with no pattern cannot go backwards --- CvxLean/Tactic/Convexify/RewriteMapLibrary.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index 76f0d8ac..afaf38f5 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -236,7 +236,7 @@ register_rewrite_map "pow_half_two"; "(pow (pow ?a 0.5) 2)" => "?a" := simp_or_rw [Real.pow_half_two (by positivity_ext)]; register_rewrite_map "pow_half_two-rev"; "?a" => "(pow (pow ?a 0.5) 2)" := - simp_or_rw [←Real.pow_half_two (by positivity_ext)]; + simp_or_rw [Real.pow_half_two (by positivity_ext)]; /- Exponential and logarithm rules. -/ From 8c24ee5e05950d12aea42cb347e4277878ac5018 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 14:12:04 -0500 Subject: [PATCH 045/189] feat: add `constant_fold` rewrite map --- CvxLean/Tactic/Convexify/RewriteMapLibrary.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index afaf38f5..af281cd2 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -301,4 +301,10 @@ register_rewrite_map "geo_folding"; "(sqrt (mul ?a ?b))" => "(geo ?a ?b)" := register_rewrite_map "norm2_folding"; "(sqrt (add (pow ?a 2) (pow ?b 2)))" => "(norm2 ?a ?b)" := rfl; + +/- Constant folding. -/ + +register_rewrite_map "constant_fold"; "?" => "?" := + norm_num; + end CvxLean From e4f38ecff5c8d29af43418131c004bc27e138756 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 14:12:17 -0500 Subject: [PATCH 046/189] fix: do not check the rule table --- rewriter/src/extract.rs | 26 ++++++++++++-------------- 1 file changed, 12 insertions(+), 14 deletions(-) diff --git a/rewriter/src/extract.rs b/rewriter/src/extract.rs index 348441db..73ea7674 100644 --- a/rewriter/src/extract.rs +++ b/rewriter/src/extract.rs @@ -36,7 +36,6 @@ pub struct Step { } fn get_step_aux( - rewrite: &Rewrite, direction: Direction, current: &FlatTerm, next: &FlatTerm, @@ -64,6 +63,7 @@ fn get_step_aux( // Out-of-context extraction. Useful for testing. let c_s = next.get_recexpr().to_string(); *expected_term = Some(c_s); + println!("ASdasd"); } } } @@ -105,7 +105,7 @@ fn get_step_aux( let children = current.children.iter().zip(next.children.iter()); for (left, right) in children { let child_res = - get_step_aux(rewrite, direction, left, right, location, expected_term); + get_step_aux(direction, left, right, location, expected_term); if child_res.is_some() { return child_res; } @@ -115,13 +115,11 @@ fn get_step_aux( return None; } -fn get_step(rule_table: &HashMap, current: &FlatTerm, next: &FlatTerm) -> Option { - if let Some((rewrite_name, direction)) = get_rewrite_name_and_direction(next) { - if let Some(rewrite) = rule_table.get(&rewrite_name) { - let location = &mut None; - let expected_term = &mut None; - return get_step_aux(rewrite, direction, current, next, location, expected_term); - } +fn get_step(current: &FlatTerm, next: &FlatTerm) -> Option { + if let Some((_rewrite_name, direction)) = get_rewrite_name_and_direction(next) { + let location = &mut None; + let expected_term = &mut None; + return get_step_aux(direction, current, next, location, expected_term); } return None; } @@ -217,17 +215,17 @@ pub fn get_steps_from_string(prob_s: &str, domains_vec: Vec<(String, Domain)>, d let flat_explanation : &FlatExplanation = explanation.make_flat_explanation(); - let rules_copy = rules().clone(); - let rule_table = make_rule_table(&rules_copy); - let mut res = Vec::new(); if best_cost.0 <= Curvature::Convex { for i in 0..flat_explanation.len() - 1 { let current = &flat_explanation[i]; let next = &flat_explanation[i + 1]; - match get_step(&rule_table, current, next) { + match get_step(current, next) { Some(step) => { res.push(step); } - None => { } + None => { + // Should not get here. + println!("Failed to extract step."); + } } } } else { From 63e45f2e1efbd1f6f18996d7003936f867973d91 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 14:13:34 -0500 Subject: [PATCH 047/189] wip: constant repr approach --- rewriter/src/optimization.rs | 120 +++++++++++++++++++++++++++++------ 1 file changed, 100 insertions(+), 20 deletions(-) diff --git a/rewriter/src/optimization.rs b/rewriter/src/optimization.rs index 8b63cc75..e5575b11 100644 --- a/rewriter/src/optimization.rs +++ b/rewriter/src/optimization.rs @@ -49,6 +49,42 @@ pub struct Data { pub is_constant: bool, } +fn make_constant_repr_unary(symbol: &str, f: F, val_o: Option) -> + Option<(Constant, PatternAst)> + where + F: Fn(Constant) -> Constant { + match val_o { + Some(val) => { + let res = f(val); + let res_f = res.into_inner(); + if ((2.0 * res_f) as u32) as f64 == (2.0 * res_f) { + Some((res, format!("({} {})", symbol, val).parse().unwrap())) + } else { + None + } + } + _ => None + } +} + +fn make_constant_repr_binary(symbol: &str, f: F, val1_o: Option, val2_o: Option) -> + Option<(Constant, PatternAst)> + where + F: Fn(Constant, Constant) -> Constant { + match (val1_o, val2_o) { + (Some(val1), Some(val2)) => { + let res = f(val1, val2); + let res_f = res.into_inner(); + if ((2.0 * res_f) as u32) as f64 == (2.0 * res_f) { + Some((res, format!("({} {} {})", symbol, val1, val2).parse().unwrap())) + } else { + None + } + } + _ => None + } +} + impl Analysis for Meta { type Data = Data; @@ -61,6 +97,7 @@ impl Analysis for Meta { let inter = d_to.intersection(&d_from); if Domain::is_empty(&inter) { // Should never get here. + println!("Empty domain."); to.domain = None } else { to.domain = Some(inter); @@ -97,17 +134,24 @@ impl Analysis for Meta { |i: &Id| egraph[*i].data.domain.clone(); let get_is_constant = |i: &Id| egraph[*i].data.is_constant.clone(); + let repr = + |i: &Id| egraph[*i].data.constant_repr.clone().map(|d| d.0); let domains_map = egraph.analysis.domains.clone(); let mut domain = None; let mut is_constant = false; + let mut constant_repr = None; match enode { Optimization::Neg(a) => { domain = domain::option_neg(get_domain(a)); is_constant = get_is_constant(a); + if is_constant { + constant_repr = + make_constant_repr_unary("neg", |x| -x, repr(a)) + } } Optimization::Sqrt(a) => { domain = domain::option_sqrt(get_domain(a)); @@ -133,22 +177,42 @@ impl Analysis for Meta { domain = domain::option_add( get_domain(a), get_domain(b)); is_constant = get_is_constant(a) && get_is_constant(b); + if is_constant { + constant_repr = + make_constant_repr_binary("add", |x, y| x + y, repr(a), repr(b)) + } } Optimization::Sub([a, b]) => { domain = domain::option_sub(get_domain(a), get_domain(b)); is_constant = get_is_constant(a) && get_is_constant(b); + if is_constant { + constant_repr = + make_constant_repr_binary("sub", |x, y| x - y, repr(a), repr(b)) + } } Optimization::Mul([a, b]) => { domain = domain::option_mul(get_domain(a), get_domain(b)); is_constant = get_is_constant(a) && get_is_constant(b); + if is_constant { + constant_repr = + make_constant_repr_binary("mul", |x, y| x * y, repr(a), repr(b)) + } } Optimization::Div([a, b]) => { domain = domain::option_div(get_domain(a), get_domain(b)); is_constant = get_is_constant(a) && get_is_constant(b); + if is_constant { + constant_repr = + make_constant_repr_binary("div", |x, y| x / y, repr(a), repr(b)) + } } Optimization::Pow([a, b]) => { domain = domain::option_pow(get_domain(a), get_domain(b)); is_constant = get_is_constant(a) && get_is_constant(b); + if is_constant { + constant_repr = + make_constant_repr_binary("pow", |x, y| Pow::pow(x, y), repr(a), repr(b)) + } } Optimization::QOL([a, b]) => { domain = domain::option_quad_over_lin(get_domain(a), get_domain(b)); @@ -186,42 +250,58 @@ impl Analysis for Meta { } _ => {} } - // NOTE: parameters are treated as constants. + // NOTE(RFM): parameters are treated as constants. is_constant = true; } Optimization::Symbol(_) => {} Optimization::Constant(f) => { domain = Some(Domain::make_singleton((*f).into_inner())); is_constant = true; + constant_repr = Some((*f, format!("{}", f).parse().unwrap())); } _ => {} } - Data { domain, is_constant } + Data { domain, is_constant, constant_repr } } fn modify(egraph: &mut egg::EGraph, id: Id) { let data = egraph[id].data.clone(); - if data.is_constant { - match data.domain { - Some(d) => { - match d.get_constant() { - Some(c) => { - // Only fold if constant is .0 or .5. - if ((2.0 * c) as u32) as f64 == (2.0 * c) { - let nn_c = NotNan::new(c).unwrap(); - let node = Optimization::Constant(nn_c); - let added = egraph.add(node); - egraph.union(id, added); - - egraph[id].assert_unique_leaves(); - } - } - _ => {} - } + if !data.is_constant { + return; + } + let c_from_domain; + match data.domain { + Some(d) => { + match d.get_constant() { + Some(c) => { c_from_domain = c; } + _ => { return; } } - _ => {} } + _ => { return; } + } + match data.constant_repr { + Some((c, pat)) => { + let c_f = c.into_inner(); + if c_f != c_from_domain { + // Should never get here. + print!("Constants do not match {c_f} {c_from_domain}."); + return; + } + // egraph.union_instantiations( + // &pat, + // &format!("{}", c_from_domain).parse().unwrap(), + // &Default::default(), + // "constant_fold".to_string(), + // ); + let nn_c = NotNan::new(c_f).unwrap(); + let node = Optimization::Constant(nn_c); + let added = egraph.add(node); + egraph.union_trusted(id, added, "constant_fold"); + + egraph[id].assert_unique_leaves(); +} + _ => {} } } } From df54fa4a1b99d78591a1a6a3a9fbabeffa1c1ecb Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 14:16:34 -0500 Subject: [PATCH 048/189] feat: simplify constant folding --- rewriter/src/extract.rs | 1 - rewriter/src/optimization.rs | 103 ++++------------------------------- rewriter/src/tests.rs | 7 +++ 3 files changed, 18 insertions(+), 93 deletions(-) diff --git a/rewriter/src/extract.rs b/rewriter/src/extract.rs index 73ea7674..45199f1d 100644 --- a/rewriter/src/extract.rs +++ b/rewriter/src/extract.rs @@ -63,7 +63,6 @@ fn get_step_aux( // Out-of-context extraction. Useful for testing. let c_s = next.get_recexpr().to_string(); *expected_term = Some(c_s); - println!("ASdasd"); } } } diff --git a/rewriter/src/optimization.rs b/rewriter/src/optimization.rs index e5575b11..b3c0fe89 100644 --- a/rewriter/src/optimization.rs +++ b/rewriter/src/optimization.rs @@ -49,42 +49,6 @@ pub struct Data { pub is_constant: bool, } -fn make_constant_repr_unary(symbol: &str, f: F, val_o: Option) -> - Option<(Constant, PatternAst)> - where - F: Fn(Constant) -> Constant { - match val_o { - Some(val) => { - let res = f(val); - let res_f = res.into_inner(); - if ((2.0 * res_f) as u32) as f64 == (2.0 * res_f) { - Some((res, format!("({} {})", symbol, val).parse().unwrap())) - } else { - None - } - } - _ => None - } -} - -fn make_constant_repr_binary(symbol: &str, f: F, val1_o: Option, val2_o: Option) -> - Option<(Constant, PatternAst)> - where - F: Fn(Constant, Constant) -> Constant { - match (val1_o, val2_o) { - (Some(val1), Some(val2)) => { - let res = f(val1, val2); - let res_f = res.into_inner(); - if ((2.0 * res_f) as u32) as f64 == (2.0 * res_f) { - Some((res, format!("({} {} {})", symbol, val1, val2).parse().unwrap())) - } else { - None - } - } - _ => None - } -} - impl Analysis for Meta { type Data = Data; @@ -134,24 +98,17 @@ impl Analysis for Meta { |i: &Id| egraph[*i].data.domain.clone(); let get_is_constant = |i: &Id| egraph[*i].data.is_constant.clone(); - let repr = - |i: &Id| egraph[*i].data.constant_repr.clone().map(|d| d.0); let domains_map = egraph.analysis.domains.clone(); let mut domain = None; let mut is_constant = false; - let mut constant_repr = None; match enode { Optimization::Neg(a) => { domain = domain::option_neg(get_domain(a)); is_constant = get_is_constant(a); - if is_constant { - constant_repr = - make_constant_repr_unary("neg", |x| -x, repr(a)) - } } Optimization::Sqrt(a) => { domain = domain::option_sqrt(get_domain(a)); @@ -177,42 +134,22 @@ impl Analysis for Meta { domain = domain::option_add( get_domain(a), get_domain(b)); is_constant = get_is_constant(a) && get_is_constant(b); - if is_constant { - constant_repr = - make_constant_repr_binary("add", |x, y| x + y, repr(a), repr(b)) - } } Optimization::Sub([a, b]) => { domain = domain::option_sub(get_domain(a), get_domain(b)); is_constant = get_is_constant(a) && get_is_constant(b); - if is_constant { - constant_repr = - make_constant_repr_binary("sub", |x, y| x - y, repr(a), repr(b)) - } } Optimization::Mul([a, b]) => { domain = domain::option_mul(get_domain(a), get_domain(b)); is_constant = get_is_constant(a) && get_is_constant(b); - if is_constant { - constant_repr = - make_constant_repr_binary("mul", |x, y| x * y, repr(a), repr(b)) - } } Optimization::Div([a, b]) => { domain = domain::option_div(get_domain(a), get_domain(b)); is_constant = get_is_constant(a) && get_is_constant(b); - if is_constant { - constant_repr = - make_constant_repr_binary("div", |x, y| x / y, repr(a), repr(b)) - } } Optimization::Pow([a, b]) => { domain = domain::option_pow(get_domain(a), get_domain(b)); is_constant = get_is_constant(a) && get_is_constant(b); - if is_constant { - constant_repr = - make_constant_repr_binary("pow", |x, y| Pow::pow(x, y), repr(a), repr(b)) - } } Optimization::QOL([a, b]) => { domain = domain::option_quad_over_lin(get_domain(a), get_domain(b)); @@ -257,12 +194,11 @@ impl Analysis for Meta { Optimization::Constant(f) => { domain = Some(Domain::make_singleton((*f).into_inner())); is_constant = true; - constant_repr = Some((*f, format!("{}", f).parse().unwrap())); } _ => {} } - Data { domain, is_constant, constant_repr } + Data { domain, is_constant } } fn modify(egraph: &mut egg::EGraph, id: Id) { @@ -270,38 +206,21 @@ impl Analysis for Meta { if !data.is_constant { return; } - let c_from_domain; match data.domain { Some(d) => { match d.get_constant() { - Some(c) => { c_from_domain = c; } - _ => { return; } + Some(c) => { + let nn_c = NotNan::new(c).unwrap(); + let node = Optimization::Constant(nn_c); + let added = egraph.add(node); + egraph.union_trusted(id, added, "constant_fold"); + + egraph[id].assert_unique_leaves(); + } + _ => { } } } - _ => { return; } - } - match data.constant_repr { - Some((c, pat)) => { - let c_f = c.into_inner(); - if c_f != c_from_domain { - // Should never get here. - print!("Constants do not match {c_f} {c_from_domain}."); - return; - } - // egraph.union_instantiations( - // &pat, - // &format!("{}", c_from_domain).parse().unwrap(), - // &Default::default(), - // "constant_fold".to_string(), - // ); - let nn_c = NotNan::new(c_f).unwrap(); - let node = Optimization::Constant(nn_c); - let added = egraph.add(node); - egraph.union_trusted(id, added, "constant_fold"); - - egraph[id].assert_unique_leaves(); -} - _ => {} + _ => { } } } } diff --git a/rewriter/src/tests.rs b/rewriter/src/tests.rs index aeb19b84..0363f1b5 100644 --- a/rewriter/src/tests.rs +++ b/rewriter/src/tests.rs @@ -309,4 +309,11 @@ fn test_3_33() { ); } +// Constant folding. + +#[test] +fn test_constant_folding() { + assert_steps_from_string("(add 2 2)"); +} + } From 8573e4e26bef740dda55b3d9ff598c81a57451e2 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 14:21:50 -0500 Subject: [PATCH 049/189] chore: deal with unused code --- rewriter/src/explain_utils.rs | 1 + rewriter/src/extract.rs | 3 --- 2 files changed, 1 insertion(+), 3 deletions(-) diff --git a/rewriter/src/explain_utils.rs b/rewriter/src/explain_utils.rs index c83b3b41..9a260658 100644 --- a/rewriter/src/explain_utils.rs +++ b/rewriter/src/explain_utils.rs @@ -29,6 +29,7 @@ pub fn get_rewrite_name_and_direction(t : &FlatTerm) -> Option<(Sy return None; } +#[allow(unused)] pub fn make_rule_table(rules: &Vec) -> HashMap { let mut table: HashMap = Default::default(); for r in rules { diff --git a/rewriter/src/extract.rs b/rewriter/src/extract.rs index 45199f1d..7a032efc 100644 --- a/rewriter/src/extract.rs +++ b/rewriter/src/extract.rs @@ -22,11 +22,8 @@ use cost::DCPCost as DCPCost; use crate::explain_utils; use explain_utils::Direction as Direction; -use explain_utils::make_rule_table as make_rule_table; use explain_utils::get_rewrite_name_and_direction as get_rewrite_name_and_direction; -pub type Rewrite = egg::Rewrite; - #[derive(Serialize, Debug)] pub struct Step { rewrite_name : String, From 3360c4cade981479bbb2eced3baa58bc41eb8b4a Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 14:28:54 -0500 Subject: [PATCH 050/189] wip: turns out constant folding is not needed, comment out for now --- CvxLean/Tactic/Convexify/RewriteMapLibrary.lean | 6 ------ CvxLean/Test/Convexify/Unit.lean | 2 +- 2 files changed, 1 insertion(+), 7 deletions(-) diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index af281cd2..afaf38f5 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -301,10 +301,4 @@ register_rewrite_map "geo_folding"; "(sqrt (mul ?a ?b))" => "(geo ?a ?b)" := register_rewrite_map "norm2_folding"; "(sqrt (add (pow ?a 2) (pow ?b 2)))" => "(norm2 ?a ?b)" := rfl; - -/- Constant folding. -/ - -register_rewrite_map "constant_fold"; "?" => "?" := - norm_num; - end CvxLean diff --git a/CvxLean/Test/Convexify/Unit.lean b/CvxLean/Test/Convexify/Unit.lean index d32aa8cf..2a34584f 100644 --- a/CvxLean/Test/Convexify/Unit.lean +++ b/CvxLean/Test/Convexify/Unit.lean @@ -388,7 +388,7 @@ time_cmd equivalence mulZeroConstrRed/mulZeroConstrAuto : mulZeroConstr := by #print mulZeroConstrAuto -- mul_comm (obj) - +-- NOTE(RFM): Empty domain here. Why? def mulCommObj := optimization (x y : ℝ) minimize (x * (y * (1 / x)) : ℝ) From 475e4495eed52f3016279de11fb90e88c72b7b3d Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 15:39:58 -0500 Subject: [PATCH 051/189] chore: remove comments in power atom --- CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean index 78097d32..366327e2 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean @@ -10,7 +10,6 @@ namespace CvxLean open Real -set_option trace.Meta.debug true in declare_atom powNegOne [convex] (x : ℝ)- : x ^ (-1) := vconditions (hx : 0 < x) @@ -19,7 +18,7 @@ implementationObjective (t) implementationConstraints (c1 : soCone (t + x) ![t - x, 2]) (c2 : 0 ≤ t) - (c3 : 0 ≤ x) -- TODO: This is not reduced. + (c3 : 0 ≤ x) solution (t := x ^ (-1)) solutionEqualsAtom rfl @@ -66,7 +65,7 @@ implementationConstraints (c2 : soCone (t₀ + 1) ![t₀ - 1, 2 * t₁]) (c3 : 0 ≤ t₀) (c4 : 0 ≤ t₁) - (c5 : 0 ≤ x) -- TODO: This is not reduced. + (c5 : 0 ≤ x) solution (t₀ := x ^ (-2)) (t₁ := x ^ (-1)) solutionEqualsAtom rfl From 922e6e7ad91d5be02b21cf9db3158a53aa2e962e Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 16:23:14 -0500 Subject: [PATCH 052/189] wip: `quadOverLin` atom in Lean --- .../DCP/AtomLibrary/Fns/QuadOverLin.lean | 55 +++++++++++++++++++ 1 file changed, 55 insertions(+) create mode 100644 CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean new file mode 100644 index 00000000..7151e6ae --- /dev/null +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean @@ -0,0 +1,55 @@ +import CvxLean.Tactic.DCP.Atoms +import CvxLean.Tactic.DCP.AtomLibrary.Sets.Cones +import CvxLean.Tactic.DCP.AtomLibrary.Fns.VecCons +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Add +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Sub +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Mul +import CvxLean.Lib.Math.Data.Vec + +namespace CvxLean + +open Real + +#check sqrt_le_iff + +-- TODO: generalize to x : Fin n → ℝ +-- TODO: distinguish between nonincreasing and nondecreasing. +declare_atom quadOverLin [convex] (x : ℝ)? (y : ℝ)- : x ^ 2 / y := +vconditions + (hy : 0 < y) +implementationVars (t : ℝ) +implementationObjective (t) +implementationConstraints + (c : soCone (y + t) ![y - t, 2 * x]) +solution (t := x ^ 2 / y) +solutionEqualsAtom by rfl +feasibility + (c : by + have hynn := le_of_lt hy + have hx2ynn : 0 ≤ x ^ 2 / y := by + rw [rpow_two] + exact div_nonneg (pow_two_nonneg x) hynn + rw [soCone_add_sub_two_mul_of_nonneg _ hynn hx2ynn] + unfold soCone + simp [sqrt_le_iff] + split_ands + { apply add_nonneg (le_of_lt hy) + exact div_nonneg (pow_two_nonneg x) (le_of_lt hy) } + { have hlhs : (y - x ^ 2 / y) ^ 2 + (2 * x) ^ 2 + = (y ^ 2 + (x ^ 2 / y) ^ 2) + (-(2 * x ^ 2 * (y / y)) + 4 * x ^ 2) := by + norm_cast; ring + have hrhs : (y + x ^ 2 / y) ^ 2 + = (y ^ 2 + (x ^ 2 / y) ^ 2) + (2 * x ^ 2 * (y / y)) := by + norm_cast; ring + simp only [←rpow_two] + rw [hlhs, hrhs, div_self (ne_of_gt hy)] + apply add_le_add (le_refl _) + ring_nf! + exact le_refl _ }) +optimality by + intros y' hyy' + sorry +vconditionElimination + (hy : sorry) + +end CvxLean From 7e0aa15442ed2f3a28094b03c7127581f3088e80 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 16:23:38 -0500 Subject: [PATCH 053/189] feat: another lemma about `soCone` --- CvxLean/Lib/Cones/SOCone.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/CvxLean/Lib/Cones/SOCone.lean b/CvxLean/Lib/Cones/SOCone.lean index 81c22717..cb44e0b4 100644 --- a/CvxLean/Lib/Cones/SOCone.lean +++ b/CvxLean/Lib/Cones/SOCone.lean @@ -49,12 +49,21 @@ lemma soCone_add_sub_two_mul_of_nonneg {x y : ℝ} (z : ℝ) (hx : 0 ≤ x) (hy conv => lhs; unfold soCone; simp [sqrt_le_left hxy, ←le_sub_iff_add_le'] ring_nf; simp +/-- Same as `soCone_add_sub_two_mul_of_nonneg` with z = 1. -/ lemma soCone_add_sub_two_of_nonneg {x y : ℝ} (hx : 0 ≤ x) (hy : 0 ≤ y) : soCone (x + y) ![x - y, 2] ↔ 1 ≤ x * y := by have h := soCone_add_sub_two_mul_of_nonneg 1 hx hy rw [mul_one, one_rpow] at h exact h +/-- Similar trick. -/ +lemma soCone_sub_add_two_mul_of_nonneg {x y : ℝ} (z : ℝ) : + soCone (x - y) ![x + y, 2 * z] ↔ y ≤ x ∧ z ^ (2 : ℝ) ≤ -(x * y) := by + conv => lhs; unfold soCone; simp [sqrt_le_iff, ←le_sub_iff_add_le'] + apply Iff.and + { rfl } + { ring_nf!; rw [←neg_mul, ←div_le_iff (by norm_num)]; simp } + end Lemmas end Real From 760f7eeb6b257e33caf85fe6349ba70ae2aa9e14 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 16:43:19 -0500 Subject: [PATCH 054/189] feat: graph implementation for `quadOverLin` --- .../DCP/AtomLibrary/Fns/QuadOverLin.lean | 60 ++++++++++--------- 1 file changed, 31 insertions(+), 29 deletions(-) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean index 7151e6ae..3c8e33cb 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean @@ -1,55 +1,57 @@ import CvxLean.Tactic.DCP.Atoms import CvxLean.Tactic.DCP.AtomLibrary.Sets.Cones +import CvxLean.Tactic.DCP.AtomLibrary.Sets.Le import CvxLean.Tactic.DCP.AtomLibrary.Fns.VecCons import CvxLean.Tactic.DCP.AtomLibrary.Fns.Add import CvxLean.Tactic.DCP.AtomLibrary.Fns.Sub import CvxLean.Tactic.DCP.AtomLibrary.Fns.Mul +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Exp import CvxLean.Lib.Math.Data.Vec namespace CvxLean open Real -#check sqrt_le_iff - --- TODO: generalize to x : Fin n → ℝ --- TODO: distinguish between nonincreasing and nondecreasing. +-- TODO(RFM): generalize to x : Fin n → ℝ +-- TODO(RFM): distinguish between nonincreasing and nondecreasing. declare_atom quadOverLin [convex] (x : ℝ)? (y : ℝ)- : x ^ 2 / y := vconditions (hy : 0 < y) -implementationVars (t : ℝ) +implementationVars (t : ℝ) (y' : ℝ) implementationObjective (t) implementationConstraints - (c : soCone (y + t) ![y - t, 2 * x]) -solution (t := x ^ 2 / y) + (c1 : soCone (y + t) ![y - t, 2 * x]) + (c2 : 0 ≤ t) + -- NOTE(RFM): This is a trick to make y strictly positive. + (c3 : exp y' ≤ y) +solution (t := x ^ 2 / y) (y' := log y) solutionEqualsAtom by rfl feasibility - (c : by + (c1 : by have hynn := le_of_lt hy have hx2ynn : 0 ≤ x ^ 2 / y := by - rw [rpow_two] - exact div_nonneg (pow_two_nonneg x) hynn + rw [rpow_two]; exact div_nonneg (pow_two_nonneg x) hynn rw [soCone_add_sub_two_mul_of_nonneg _ hynn hx2ynn] - unfold soCone - simp [sqrt_le_iff] - split_ands - { apply add_nonneg (le_of_lt hy) - exact div_nonneg (pow_two_nonneg x) (le_of_lt hy) } - { have hlhs : (y - x ^ 2 / y) ^ 2 + (2 * x) ^ 2 - = (y ^ 2 + (x ^ 2 / y) ^ 2) + (-(2 * x ^ 2 * (y / y)) + 4 * x ^ 2) := by - norm_cast; ring - have hrhs : (y + x ^ 2 / y) ^ 2 - = (y ^ 2 + (x ^ 2 / y) ^ 2) + (2 * x ^ 2 * (y / y)) := by - norm_cast; ring - simp only [←rpow_two] - rw [hlhs, hrhs, div_self (ne_of_gt hy)] - apply add_le_add (le_refl _) - ring_nf! - exact le_refl _ }) + rw [mul_div, mul_comm, ←mul_div, div_self (ne_of_gt hy), mul_one]) + (c2 : by + have hynn := le_of_lt hy + rw [rpow_two] + exact div_nonneg (pow_two_nonneg x) hynn) + (c3 : by + simp [exp_log hy]) optimality by - intros y' hyy' - sorry + intros z hyz + have hy := lt_of_lt_of_le (exp_pos _) c3 + have hynn := le_of_lt hy + rw [soCone_add_sub_two_mul_of_nonneg x hynn c2] at c1 + have hz := lt_of_lt_of_le hy hyz + rw [div_le_iff' hz] + apply le_trans c1 + apply mul_le_mul hyz (le_refl _) c2 (le_of_lt hz) vconditionElimination - (hy : sorry) + (hy : by + intros z hz + have hy := lt_of_lt_of_le (exp_pos _) c3 + exact lt_of_lt_of_le hy hz) end CvxLean From d578bb1c73d699a74e684209921070441c61e712 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 16:49:43 -0500 Subject: [PATCH 055/189] fix: no constant folding in egg --- rewriter/src/optimization.rs | 23 ----------------------- 1 file changed, 23 deletions(-) diff --git a/rewriter/src/optimization.rs b/rewriter/src/optimization.rs index b3c0fe89..b1844791 100644 --- a/rewriter/src/optimization.rs +++ b/rewriter/src/optimization.rs @@ -200,29 +200,6 @@ impl Analysis for Meta { Data { domain, is_constant } } - - fn modify(egraph: &mut egg::EGraph, id: Id) { - let data = egraph[id].data.clone(); - if !data.is_constant { - return; - } - match data.domain { - Some(d) => { - match d.get_constant() { - Some(c) => { - let nn_c = NotNan::new(c).unwrap(); - let node = Optimization::Constant(nn_c); - let added = egraph.add(node); - egraph.union_trusted(id, added, "constant_fold"); - - egraph[id].assert_unique_leaves(); - } - _ => { } - } - } - _ => { } - } - } } pub fn is_gt_zero(var: &str) -> impl Fn(&mut EGraph, Id, &Subst) -> bool { From 85a2db683414e44378cff5ac3eaaec1adf4a1794 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 17:00:56 -0500 Subject: [PATCH 056/189] feat: graph implementation for `geoMean` --- .../Tactic/DCP/AtomLibrary/Fns/GeoMean.lean | 52 +++++++++++++++++++ 1 file changed, 52 insertions(+) create mode 100644 CvxLean/Tactic/DCP/AtomLibrary/Fns/GeoMean.lean diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/GeoMean.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/GeoMean.lean new file mode 100644 index 00000000..66ba3710 --- /dev/null +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/GeoMean.lean @@ -0,0 +1,52 @@ +import CvxLean.Tactic.DCP.Atoms +import CvxLean.Tactic.DCP.AtomLibrary.Sets.Cones +import CvxLean.Tactic.DCP.AtomLibrary.Sets.Le +import CvxLean.Tactic.DCP.AtomLibrary.Fns.VecCons +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Add +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Sub +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Mul +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Exp +import CvxLean.Lib.Math.Data.Vec + +namespace CvxLean + +open Real + +-- TODO(RFM): generalize to x : Fin n → ℝ +declare_atom geoMean [concave] (x : ℝ)+ (y : ℝ)+ : sqrt (x * y) := +vconditions + (hx : 0 ≤ x) + (hy : 0 ≤ y) +implementationVars (t : ℝ) +implementationObjective (t) +implementationConstraints + (c1 : soCone (x + y) ![x - y, 2 * t]) + (c2 : 0 ≤ t) + (c3 : 0 ≤ x) + (c4 : 0 ≤ y) +solution (t := sqrt (x * y)) +solutionEqualsAtom by rfl +feasibility + (c1 : by + rw [soCone_add_sub_two_mul_of_nonneg _ hx hy, rpow_two, sq_sqrt] + exact mul_nonneg hx hy) + (c2 : by + exact sqrt_nonneg _) + (c3 : hx) + (c4 : hy) +optimality by + intros v w hxv hyw + have hv := le_trans c3 hxv + have hw := le_trans c4 hyw + rw [soCone_add_sub_two_mul_of_nonneg _ c3 c4, rpow_two] at c1 + rw [le_sqrt c2 (mul_nonneg hv hw)] + exact le_trans c1 (mul_le_mul hxv hyw c4 hv) +vconditionElimination + (hx : by + intros v _ hxv _ + exact le_trans c3 hxv) + (hy : by + intros _ w _ hyw + exact le_trans c4 hyw) + +end CvxLean From caf5beaf4a359e6969524859ff66750946e75a9d Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 17:10:54 -0500 Subject: [PATCH 057/189] test: newly added atoms --- CvxLean/Tactic/DCP/AtomLibrary/Fns/All.lean | 2 + CvxLean/Test/DCP/Dcp.lean | 48 +++++++++++++++++++++ 2 files changed, 50 insertions(+) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/All.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/All.lean index fb9cd0a3..31fa3c72 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/All.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/All.lean @@ -7,6 +7,7 @@ import CvxLean.Tactic.DCP.AtomLibrary.Fns.DotProduct import CvxLean.Tactic.DCP.AtomLibrary.Fns.Entr import CvxLean.Tactic.DCP.AtomLibrary.Fns.Exp import CvxLean.Tactic.DCP.AtomLibrary.Fns.FromBlocks +import CvxLean.Tactic.DCP.AtomLibrary.Fns.GeoMean import CvxLean.Tactic.DCP.AtomLibrary.Fns.Huber import CvxLean.Tactic.DCP.AtomLibrary.Fns.KLDiv import CvxLean.Tactic.DCP.AtomLibrary.Fns.Log @@ -21,6 +22,7 @@ import CvxLean.Tactic.DCP.AtomLibrary.Fns.Nth import CvxLean.Tactic.DCP.AtomLibrary.Fns.Perspective import CvxLean.Tactic.DCP.AtomLibrary.Fns.PosSemidef import CvxLean.Tactic.DCP.AtomLibrary.Fns.Power +import CvxLean.Tactic.DCP.AtomLibrary.Fns.QuadOverLin import CvxLean.Tactic.DCP.AtomLibrary.Fns.SMul import CvxLean.Tactic.DCP.AtomLibrary.Fns.Sq import CvxLean.Tactic.DCP.AtomLibrary.Fns.Sqrt diff --git a/CvxLean/Test/DCP/Dcp.lean b/CvxLean/Test/DCP/Dcp.lean index e1599c9a..cadf9de0 100644 --- a/CvxLean/Test/DCP/Dcp.lean +++ b/CvxLean/Test/DCP/Dcp.lean @@ -174,3 +174,51 @@ noncomputable example : Solution $ hsqrt : 0 ≤ x := by dcp -- TODO: first constraint not reduced. We need to handle constant constraints. sorry + +noncomputable def testXExp : Solution $ + optimization (x : ℝ) + minimize x * exp x + subject to + c0 : 0 ≤ x +:= by + dcp + sorry + +noncomputable def testEntr : Solution $ + optimization (x : ℝ) + minimize -(-(x * log x)) + subject to + c0 : 0 ≤ x +:= by + dcp + sorry + +noncomputable def testQuadOverLin : Solution $ + optimization (x y : ℝ) + minimize x ^ 2 / y + subject to + c0 : 0 ≤ x + c0 : 0.001 ≤ y +:= by + dcp + sorry + +noncomputable def testGeoMean : Solution $ + optimization (x y : ℝ) + minimize - (sqrt (x * y)) + subject to + c0 : 0 ≤ x + c0 : 0 ≤ y +:= by + dcp + sorry + +noncomputable def testNorm2 : Solution $ + optimization (x y : ℝ) + minimize ‖![x, y]‖ + subject to + c0 : 0 ≤ x + c0 : 0 ≤ y +:= by + dcp + sorry From 404ddc08e1373accdc0c2672169ec97ecafc19ea Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 17:28:16 -0500 Subject: [PATCH 058/189] =?UTF-8?q?feat:=20`norm2=E2=82=82`=20atom=20for?= =?UTF-8?q?=20pattern-matching?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- CvxLean/Tactic/DCP/AtomLibrary/Fns/Norm.lean | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Norm.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Norm.lean index b41df638..94ee4cb3 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Norm.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Norm.lean @@ -1,5 +1,6 @@ import CvxLean.Tactic.DCP.Atoms import CvxLean.Tactic.DCP.AtomLibrary.Sets.Cones +import CvxLean.Tactic.DCP.AtomLibrary.Fns.VecCons import CvxLean.Lib.Math.Data.Vec namespace CvxLean @@ -24,4 +25,22 @@ optimality by exact c vconditionElimination +declare_atom norm2₂ [convex] (x : ℝ)? (y : ℝ)? : sqrt (x ^ 2 + y ^ 2) := +vconditions +implementationVars (t : ℝ) +implementationObjective (t) +implementationConstraints + (c : soCone t ![x, y]) +solution (t := sqrt (x ^ 2 + y ^ 2)) +solutionEqualsAtom by rfl +feasibility + (c : by simp [soCone]) +optimality by + simp [soCone] at c + simp [c] +vconditionElimination + +lemma norm2₂_eq_norm2 {x y : ℝ} : ‖![x, y]‖ = sqrt (x ^ 2 + y ^ 2) := + by simp [Norm.norm] + end CvxLean From 90b6bd480c358b59a1f371a563225911a8d96361 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 17:29:57 -0500 Subject: [PATCH 059/189] feat: new atoms in egg to Lean reconstruction phase --- CvxLean/Tactic/Convexify/Egg/ToExpr.lean | 35 ++++++++++++++++++------ 1 file changed, 26 insertions(+), 9 deletions(-) diff --git a/CvxLean/Tactic/Convexify/Egg/ToExpr.lean b/CvxLean/Tactic/Convexify/Egg/ToExpr.lean index b5bca3f1..26be06de 100644 --- a/CvxLean/Tactic/Convexify/Egg/ToExpr.lean +++ b/CvxLean/Tactic/Convexify/Egg/ToExpr.lean @@ -1,4 +1,4 @@ -import CvxLean.Lib.Math.Data.Real +]import CvxLean.Lib.Math.Data.Real import CvxLean.Meta.Minimization import CvxLean.Tactic.Convexify.Egg.Sexp import CvxLean.Tactic.DCP.Tree @@ -85,6 +85,20 @@ partial def EggTree.toExpr (vars : List String) : Tree String String → MetaM E | Tree.node "sqrt" #[t] => do let t ← toExpr vars t return mkAppN (mkConst ``Real.sqrt) #[t] + -- Log. + | Tree.node "log" #[t] => do + let t ← toExpr vars t + return mkAppN (mkConst ``Real.log) #[t] + -- Exp. + | Tree.node "exp" #[t] => do + let t ← toExpr vars t + return mkAppN (mkConst ``Real.exp) #[t] + -- XExp. + | Tree.node "xexp" #[t] => + EggTree.toExpr vars (Tree.node "mul" #[t, Tree.node "exp" #[t]]) + -- Entr. + | Tree.node "entr" #[t] => + EggTree.toExpr vars (Tree.node "neg" #[Tree.node "mul" #[t, Tree.node "log" #[t]]]) -- Addition. | Tree.node "add" #[t1, t2] => do let t1 ← toExpr vars t1 @@ -105,19 +119,22 @@ partial def EggTree.toExpr (vars : List String) : Tree String String → MetaM E let t1 ← toExpr vars t1 let t2 ← toExpr vars t2 return mkRealHBinAppExpr ``HDiv.hDiv ``instHDiv 1 ``Real.instDivReal t1 t2 - -- Log. - | Tree.node "log" #[t] => do - let t ← toExpr vars t - return mkAppN (mkConst ``Real.log) #[t] - -- Exp. - | Tree.node "exp" #[t] => do - let t ← toExpr vars t - return mkAppN (mkConst ``Real.exp) #[t] -- Pow. | Tree.node "pow" #[t1, t2] => do let t1 ← toExpr vars t1 let t2 ← toExpr vars t2 return mkRealHBinAppExpr ``HPow.hPow ``instHPow 2 ``Real.instPowReal t1 t2 + -- Quad over Lin. + | Tree.node "qol" #[t1, t2] => + EggTree.toExpr vars (Tree.node "div" #[Tree.node "pow" #[t1, Tree.leaf "2"], t2]) + -- Geo mean. + | Tree.node "geo" #[t1, t2] => + EggTree.toExpr vars (Tree.node "sqrt" #[Tree.node "mul" #[t1, t2]]) + -- Norm2. + | Tree.node "norm2" #[t1, t2] => + EggTree.toExpr vars (Tree.node "sqrt" #[Tree.node "add" #[ + Tree.node "pow" #[t1, Tree.leaf "2"], + Tree.node "pow" #[t2, Tree.leaf "2"]]]) -- Constr. | Tree.node "constr" #[Tree.leaf s, t] => do let t ← toExpr vars t From f1770d2ce18be83d51c2edde607f7edc1198327b Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 17:52:17 -0500 Subject: [PATCH 060/189] feat: allow atom unfolding --- .../Tactic/Convexify/RewriteMapLibrary.lean | 25 ++++++++++++++---- rewriter/src/rules.rs | 26 ++++++++++++++----- 2 files changed, 40 insertions(+), 11 deletions(-) diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index afaf38f5..f011e5c1 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -286,19 +286,34 @@ register_rewrite_map "log_exp" ; "(log (exp ?a))" => "?a" := /- Atom folding. -/ -register_rewrite_map "xexp_folding"; "(mul ?a (exp ?a))" => "(xexp ?a)" := +register_rewrite_map "xexp_fold"; "(mul ?a (exp ?a))" => "(xexp ?a)" := rfl; -register_rewrite_map "entr_folding"; "(neg (mul ?a (log ?a)))" => "(entr ?a)" := +register_rewrite_map "xexp_unfold"; "(xexp ?a)" => "(mul ?a (exp ?a))" := rfl; -register_rewrite_map "qol_folding"; "(div (pow ?a 2) ?b)" => "(qol ?a ?b)" := +register_rewrite_map "entr_fold"; "(neg (mul ?a (log ?a)))" => "(entr ?a)" := rfl; -register_rewrite_map "geo_folding"; "(sqrt (mul ?a ?b))" => "(geo ?a ?b)" := +register_rewrite_map "entr_unfold"; "(entr ?a)" => "(neg (mul ?a (log ?a)))" := rfl; -register_rewrite_map "norm2_folding"; "(sqrt (add (pow ?a 2) (pow ?b 2)))" => "(norm2 ?a ?b)" := +register_rewrite_map "qol_fold"; "(div (pow ?a 2) ?b)" => "(qol ?a ?b)" := + rfl; + +register_rewrite_map "qol_unfold"; "(qol ?a ?b)" => "(div (pow ?a 2) ?b)" := + rfl; + +register_rewrite_map "geo_fold"; "(sqrt (mul ?a ?b))" => "(geo ?a ?b)" := + rfl; + +register_rewrite_map "geo_unfold"; "(geo ?a ?b)" => "(sqrt (mul ?a ?b))" := + rfl; + +register_rewrite_map "norm2_fold"; "(sqrt (add (pow ?a 2) (pow ?b 2)))" => "(norm2 ?a ?b)" := + rfl; + +register_rewrite_map "norm2_unfold"; "(norm2 ?a ?b)" => "(sqrt (add (pow ?a 2) (pow ?b 2)))" := rfl; end CvxLean diff --git a/rewriter/src/rules.rs b/rewriter/src/rules.rs index fbe798cc..5838ad16 100644 --- a/rewriter/src/rules.rs +++ b/rewriter/src/rules.rs @@ -198,21 +198,35 @@ pub fn rules() -> Vec> { vec![ rw!("log_exp"; "(log (exp ?a))" => "?a"), - /* Atom folding rules. */ + /* Atom folding and unfolding rules. */ - rw!("xexp_folding"; "(mul ?a (exp ?a))" => "(xexp ?a)" + rw!("xexp_fold"; "(mul ?a (exp ?a))" => "(xexp ?a)" + if is_ge_zero("?a")), + + rw!("xexp_unfold"; "(xexp ?a)" => "(mul ?a (exp ?a))" if is_ge_zero("?a")), - rw!("entr_folding"; "(neg (mul ?a (log ?a)))" => "(entr ?a)" + rw!("entr_fold"; "(neg (mul ?a (log ?a)))" => "(entr ?a)" + if is_gt_zero("?a")), + + rw!("entr_unfold"; "(entr ?a)" => "(neg (mul ?a (log ?a)))" if is_gt_zero("?a")), - rw!("qol_folding"; "(div (pow ?a 2) ?b)" => "(qol ?a ?b)" + rw!("qol_fold"; "(div (pow ?a 2) ?b)" => "(qol ?a ?b)" + if is_gt_zero("?b")), + + rw!("qol_unfold"; "(qol ?a ?b)" => "(div (pow ?a 2) ?b)" if is_gt_zero("?b")), - rw!("geo_folding"; "(sqrt (mul ?a ?b))" => "(geo ?a ?b)" + rw!("geo_fold"; "(sqrt (mul ?a ?b))" => "(geo ?a ?b)" + if is_gt_zero("?a") if is_gt_zero("?b")), + + rw!("geo_unfold"; "(geo ?a ?b)" => "(sqrt (mul ?a ?b))" if is_gt_zero("?a") if is_gt_zero("?b")), - rw!("norm2_folding"; "(sqrt (add (pow ?a 2) (pow ?b 2)))" => "(norm2 ?a ?b)"), + rw!("norm2_fold"; "(sqrt (add (pow ?a 2) (pow ?b 2)))" => "(norm2 ?a ?b)"), + + rw!("norm2_unfold"; "(norm2 ?a ?b)" => "(sqrt (add (pow ?a 2) (pow ?b 2)))"), ] } #[allow(unused)] From f1cb06665b9db33135e9f8c5327758a65b8c728c Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 17:52:40 -0500 Subject: [PATCH 061/189] feat: translation between Lean and egg atoms --- CvxLean/Tactic/Convexify/Egg/FromMinimization.lean | 7 ++++--- CvxLean/Tactic/Convexify/Egg/ToExpr.lean | 4 ++-- 2 files changed, 6 insertions(+), 5 deletions(-) diff --git a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean index efc04635..5e94b5b5 100644 --- a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean +++ b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean @@ -48,16 +48,17 @@ def _root_.EggTree.opMap : HashMap String (String × Nat × Array String) := ("le", ("le", 2, #[])), ("lt", ("lt", 2, #[])), ("neg", ("neg", 1, #[])), + ("sqrt'", ("sqrt", 1, #[])), + ("log", ("log", 1, #[])), + ("exp", ("exp", 1, #[])), ("add", ("add", 2, #[])), ("sub", ("sub", 2, #[])), ("mul1", ("mul", 2, #[])), ("mul2", ("mul", 2, #[])), ("sq", ("pow", 1, #["2"])), ("sqrt", ("sqrt", 1, #[])), - ("sqrt'", ("sqrt", 1, #[])), ("div", ("div", 2, #[])), - ("log", ("log", 1, #[])), - ("exp", ("exp", 1, #[])) + ("quadOverLin", ("qol", 2, #[])) ] /-- Traverse the tree and use `EggTree.opMap` to align the names of the diff --git a/CvxLean/Tactic/Convexify/Egg/ToExpr.lean b/CvxLean/Tactic/Convexify/Egg/ToExpr.lean index 26be06de..4024cc64 100644 --- a/CvxLean/Tactic/Convexify/Egg/ToExpr.lean +++ b/CvxLean/Tactic/Convexify/Egg/ToExpr.lean @@ -1,4 +1,4 @@ -]import CvxLean.Lib.Math.Data.Real +import CvxLean.Lib.Math.Data.Real import CvxLean.Meta.Minimization import CvxLean.Tactic.Convexify.Egg.Sexp import CvxLean.Tactic.DCP.Tree @@ -133,7 +133,7 @@ partial def EggTree.toExpr (vars : List String) : Tree String String → MetaM E -- Norm2. | Tree.node "norm2" #[t1, t2] => EggTree.toExpr vars (Tree.node "sqrt" #[Tree.node "add" #[ - Tree.node "pow" #[t1, Tree.leaf "2"], + Tree.node "pow" #[t1, Tree.leaf "2"], Tree.node "pow" #[t2, Tree.leaf "2"]]]) -- Constr. | Tree.node "constr" #[Tree.leaf s, t] => do From 5d8c4026e9fd5857848f58510177ff7ae22a0b83 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 17:52:58 -0500 Subject: [PATCH 062/189] wip: making convexify unit tests pass again --- CvxLean/Test/Convexify/Unit.lean | 48 +++++++++++++++++--------------- 1 file changed, 25 insertions(+), 23 deletions(-) diff --git a/CvxLean/Test/Convexify/Unit.lean b/CvxLean/Test/Convexify/Unit.lean index 2a34584f..f11fe568 100644 --- a/CvxLean/Test/Convexify/Unit.lean +++ b/CvxLean/Test/Convexify/Unit.lean @@ -13,7 +13,7 @@ open CvxLean Minimization Real /- Objective function rules. -/ -- map_objFun_log (obj) --- NOTE: This works because an affine objective is prefered. +-- NOTE(RFM): This works because an affine objective is prefered. def mapObjFunLogObj := optimization (x : ℝ) minimize (exp x) @@ -56,7 +56,7 @@ time_cmd equivalence logEqLogConstrRed/logEqLogConstrAuto : logEqLogConstr := by /- Less than or equal rules. -/ -- le_sub_iff_add_le (constr) --- NOTE: This uses le_sub_iff_add_le because 2 * x is preferred over x + x. +-- NOTE(RFM): This uses le_sub_iff_add_le because 2 * x is preferred over x + x. def leSubIffAddLeConstr := optimization (x y : ℝ) minimize (0 : ℝ) @@ -81,7 +81,7 @@ time_cmd equivalence leSubIffAddLeConstrRevRed/leSubIffAddLeConstrRevAuto : leSu #print leSubIffAddLeConstrRevAuto -- sub_le_iff_le_add (constr) --- NOTE: same reasoning as le_sub_iff_add_le. +-- NOTE(RFM): same reasoning as le_sub_iff_add_le. def subLeIffLeAddConstr := optimization (x y : ℝ) minimize (0 : ℝ) @@ -237,7 +237,7 @@ time_cmd equivalence addZeroConstrRed/addZeroConstrAuto : addZeroConstr := by #print addZeroConstrAuto -- add_comm (obj) --- NOTE: This uses one_mul-rev because 2 * x is preferred over x + x. +-- NOTE(RFM): This uses one_mul-rev because 2 * x is preferred over x + x. def addCommObj := optimization (x : ℝ) minimize (x + (1 + x) : ℝ) @@ -250,7 +250,7 @@ time_cmd equivalence addCommObjRed/addCommObjAuto : addCommObj := by #print addCommObjAuto -- add_comm (constr) --- NOTE: This uses one_mul-rev because 2 * x is preferred over x + x. +-- NOTE(RFM): This uses one_mul-rev because 2 * x is preferred over x + x. def addCommConstr := optimization (x : ℝ) minimize (0 : ℝ) @@ -264,7 +264,7 @@ time_cmd equivalence addCommConstrRed/addCommConstrAuto : addCommConstr := by #print addCommConstrAuto -- add_assoc (obj) --- NOTE: This uses one_mul-rev because 2 * x is preferred over x + x. +-- NOTE(RFM): This uses one_mul-rev because 2 * x is preferred over x + x. def addAssocObj := optimization (x : ℝ) minimize (x + (x + 1) : ℝ) @@ -277,7 +277,7 @@ time_cmd equivalence addAssocObjRed/addAssocObjAuto : addAssocObj := by #print addAssocObjAuto -- add_assoc (constr) --- NOTE: This uses one_mul-rev because 2 * x is preferred over x + x. +-- NOTE(RFM): This uses one_mul-rev because 2 * x is preferred over x + x. def addAssocConstr := optimization (x : ℝ) minimize (0 : ℝ) @@ -338,7 +338,7 @@ time_cmd equivalence oneMulConstrRed/oneMulConstrAuto : oneMulConstr := by #print oneMulConstrAuto -- one_mul-rev (obj) --- NOTE: This uses one_mul-rev because 2 * x is preferred over x + x. +-- NOTE(RFM): This uses one_mul-rev because 2 * x is preferred over x + x. def oneMulRevObj := optimization (x : ℝ) minimize (x + x : ℝ) @@ -351,7 +351,7 @@ time_cmd equivalence oneMulRevObjRed/oneMulRevObjAuto : oneMulRevObj := by #print oneMulRevObjAuto -- one_mul-rev (constr) --- NOTE: This uses one_mul-rev because 2 * x is preferred over x + x. +-- NOTE(RFM): This uses one_mul-rev because 2 * x is preferred over x + x. def oneMulRevConstr := optimization (x : ℝ) minimize (0 : ℝ) @@ -866,10 +866,10 @@ time_cmd equivalence divSelfConstrRed/divSelfConstrAuto : divSelfConstr := by /- Power and square root rules. -/ -- pow_add (obj) --- TODO: Implement power atom. +-- TODO(RFM): Implement power atom. -- pow_add (constr) --- TODO: Implement power atom. +-- TODO(RFM): Implement power atom. -- pow_add-rev (obj) def powAddRevObj := @@ -947,10 +947,10 @@ time_cmd equivalence mulPowRevConstrRed/mulPowRevConstrAuto : mulPowRevConstr := #print mulPowRevConstrAuto -- pow_mul (obj) --- TODO: Implement power atom. +-- TODO(RFM): Implement power atom. -- pow_mul (constr) --- TODO: Implement power atom. +-- TODO(RFM): Implement power atom. -- pow_mul-rev (obj) def powMulRevObj := @@ -1009,10 +1009,11 @@ def divPowRevObj := subject to hx : 0 < x -time_cmd equivalence divPowRevObjRed/divPowRevObjAuto : divPowRevObj := by - convexify +-- time_cmd equivalence divPowRevObjRed/divPowRevObjAuto : divPowRevObj := by + -- convexify +-- TODO(RFM): does not rewrite the correct term. -#print divPowRevObjAuto +-- #print divPowRevObjAuto -- div_pow-rev (constr) def divPowRevConstr := @@ -1022,16 +1023,17 @@ def divPowRevConstr := hx : 0 < x h : ((x + x) ^ 2 / x ^ 2) ≤ x -time_cmd equivalence divPowRevConstrRed/divPowRevConstrAuto : divPowRevConstr := by - convexify +-- time_cmd equivalence divPowRevConstrRed/divPowRevConstrAuto : divPowRevConstr := by +-- convexify +-- TODO(RFM): does not rewrite the correct term. -#print divPowRevConstrAuto +-- #print divPowRevConstrAuto -- pow_sub (obj) --- TODO: Implement power atom. +-- TODO(RFM): Implement power atom. -- pow_sub (constr) --- TODO: Implement power atom. +-- TODO(RFM): Implement power atom. -- pow_sub-rev (obj) def powSubRevObj := @@ -1084,10 +1086,10 @@ time_cmd equivalence divPowEqMulPowNegConstrRed/divPowEqMulPowNegConstrAuto : di #print divPowEqMulPowNegConstrAuto -- div_pow_eq_mul_pow_neg-rev (obj) --- TODO: Implement power atom. +-- TODO(RFM): Implement power atom. -- div_pow_eq_mul_pow_neg-rev (constr) --- TODO: Implement power atom. +-- TODO(RFM): Implement power atom. -- one_div_eq_pow_neg_one (obj) def oneDivEqPowNegOneObj := From e8ae2354892512cf9ba668d95ff6abe906dfa1ff Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 5 Dec 2023 18:00:58 -0500 Subject: [PATCH 063/189] wip: explanation length optimization --- rewriter/src/extract.rs | 1 + 1 file changed, 1 insertion(+) diff --git a/rewriter/src/extract.rs b/rewriter/src/extract.rs index 7a032efc..fd89bc0d 100644 --- a/rewriter/src/extract.rs +++ b/rewriter/src/extract.rs @@ -175,6 +175,7 @@ pub fn get_steps_from_string(prob_s: &str, domains_vec: Vec<(String, Domain)>, d let runner: Runner = Runner::new(analysis) .with_explanations_enabled() + .with_explanation_length_optimization() .with_node_limit(node_limit) .with_iter_limit(iter_limit) .with_time_limit(Duration::from_secs(5)) From 7ff13242050c66e972427e279b106154087760e6 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 6 Dec 2023 09:01:38 -0500 Subject: [PATCH 064/189] feat: add missing translations to `EggTree.opMap` --- CvxLean/Tactic/Convexify/Egg/FromMinimization.lean | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean index 5e94b5b5..ce661141 100644 --- a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean +++ b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean @@ -51,6 +51,8 @@ def _root_.EggTree.opMap : HashMap String (String × Nat × Array String) := ("sqrt'", ("sqrt", 1, #[])), ("log", ("log", 1, #[])), ("exp", ("exp", 1, #[])), + ("xexp", ("xexp", 1, #[])), + ("entr", ("entr", 1, #[])), ("add", ("add", 2, #[])), ("sub", ("sub", 2, #[])), ("mul1", ("mul", 2, #[])), @@ -58,7 +60,9 @@ def _root_.EggTree.opMap : HashMap String (String × Nat × Array String) := ("sq", ("pow", 1, #["2"])), ("sqrt", ("sqrt", 1, #[])), ("div", ("div", 2, #[])), - ("quadOverLin", ("qol", 2, #[])) + ("quadOverLin", ("qol", 2, #[])), + ("geoMean", ("qol", 2, #[])), + ("norm2₂", ("norm2", 2, #[])) ] /-- Traverse the tree and use `EggTree.opMap` to align the names of the From 93f507f4d2064a80f0c9b1dc3ca3b3d0ac4a2803 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 6 Dec 2023 12:20:48 -0500 Subject: [PATCH 065/189] feat: new rule `pow_two` --- CvxLean/Tactic/Convexify/RewriteMapLibrary.lean | 7 +++++++ rewriter/src/rules.rs | 4 ++++ 2 files changed, 11 insertions(+) diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index f011e5c1..673324f3 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -232,9 +232,16 @@ register_rewrite_map "sqrt_eq_rpow" ; "(sqrt ?a)" => "(pow ?a 0.5)" := register_rewrite_map "sqrt_eq_rpow-rev" ; "(pow ?a 0.5)" => "(sqrt ?a)" := simp_or_rw [←Real.sqrt_eq_rpow]; +register_rewrite_map "pow_two"; "(pow ?a 2)" => "(mul ?a ?a)" := + simp_or_rw [pow_two (M := ℝ)]; + +register_rewrite_map "pow_two-rev"; "(mul ?a ?a)" => "(pow ?a 2)" := + simp_or_rw [←pow_two (M := ℝ)]; + register_rewrite_map "pow_half_two"; "(pow (pow ?a 0.5) 2)" => "?a" := simp_or_rw [Real.pow_half_two (by positivity_ext)]; +-- TODO(RFM): Technically ← but no pattern to match on otherwise. register_rewrite_map "pow_half_two-rev"; "?a" => "(pow (pow ?a 0.5) 2)" := simp_or_rw [Real.pow_half_two (by positivity_ext)]; diff --git a/rewriter/src/rules.rs b/rewriter/src/rules.rs index 5838ad16..47f69dd7 100644 --- a/rewriter/src/rules.rs +++ b/rewriter/src/rules.rs @@ -158,6 +158,10 @@ pub fn rules() -> Vec> { vec![ rw!("sqrt_eq_rpow-rev"; "(pow ?a 0.5)" => "(sqrt ?a)"), + rw!("pow_two"; "(pow ?a 2)" => "(mul ?a ?a)"), + + rw!("pow_two-rev"; "(mul ?a ?a)" => "(pow ?a 2)"), + rw!("pow_half_two"; "(pow (pow ?a 0.5) 2)" => "?a" if is_ge_zero("?a")), rw!("pow_half_two-rev"; "?a" => "(pow (pow ?a 0.5) 2)" if is_ge_zero("?a")), From 69679d1cc6ed5e0d316cca90ea6c89d4cfa67267 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 6 Dec 2023 14:44:36 -0500 Subject: [PATCH 066/189] test: tests from Stanford course --- rewriter/src/tests.rs | 92 +++++++++++++++++++++++++++++++++++++------ 1 file changed, 81 insertions(+), 11 deletions(-) diff --git a/rewriter/src/tests.rs b/rewriter/src/tests.rs index 0363f1b5..ed39b681 100644 --- a/rewriter/src/tests.rs +++ b/rewriter/src/tests.rs @@ -278,6 +278,8 @@ fn test_norm2() { "(sqrt (add (pow (var x) 2) (pow (var y) 2)))"); } +// Misc. + #[test] fn test_norm2_with_one() { assert_steps_from_string_with_domain( @@ -286,34 +288,102 @@ fn test_norm2_with_one() { } #[test] -fn test_3_32() { +fn test_sqrt_pow4() { assert_steps_from_string_with_domain( - vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], - "(div 1 (mul (var x) (var y)))"); + vec![("x", domain::nonneg_dom())], + "(sqrt (pow (var x) 4))"); } +// From Stanford course, additional exercises. + #[test] -fn test_sqrt_pow4() { +fn test_3_32() { + // 1 / (x * y) = (x^-0.5)^2 / y + // = qol(x^-0.5, y) (this is just one possibility) assert_steps_from_string_with_domain( - vec![("x", domain::nonneg_dom())], - "(sqrt (pow (var x) 4))"); + vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], + "(div 1 (mul (var x) (var y)))"); } + + #[test] fn test_3_33() { - // NOTE(RFM): x cannot be in the free domain as otherwise we cannot convert - // x^4 to (x^2)^2. + // sqrt(1 + x^4 / y) = sqrt(1^2 + (x^2 / y)^2) + // ... = norm2(1, qol(x, y)) + // NOTE(RFM): Constant unfolding issue: x cannot be in the free domain as + // otherwise we cannot convert x^4 to (x^2)^2. assert_steps_from_string_with_domain( vec![("x", domain::nonneg_dom()), ("y", domain::pos_dom())], "(sqrt (add 1 (div (pow (var x) 4) (var y))))" ); } -// Constant folding. +#[test] +fn test_3_36_a() { + // sqrt(1 + 4x^2 + 16y^2) = sqrt((2x)^2 + (sqrt(1^2 + (4y)^2))^2) + // ... = norm2(2x, norm2(1, 4y)) + assert_steps_from_string_with_domain( + vec![("x", domain::nonneg_dom()), ("y", domain::nonneg_dom())], + "(sqrt (add 1 (add (mul 4 (pow (var x) 2)) (mul 16 (pow (var y) 2)))))"); +} + +#[test] +fn test_3_36_c() { + // log(e^(2x + 3) + e^(4y + 5)) = lse(2x + 3, 4y + 5) + assert_steps_from_string_with_domain( + vec![("x", domain::free_dom()), ("y", domain::free_dom())], + "(log (add (exp (add (mul 2 (var x)) 3)) (exp (add (mul 4 (var y)) 5))))"); +} #[test] -fn test_constant_folding() { - assert_steps_from_string("(add 2 2)"); +fn test_3_38_e() { + // (sqrt(x) + sqrt(y))^2 = x + y + 2sqrt(xy) + // ... = x + y + 2geo(x, y) + assert_steps_from_string_with_domain( + vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], + "(neg (pow (add (sqrt (var x)) (sqrt (var y))) 2))"); +} + +#[test] +fn test_3_67() { + // Generalizaiton of 3.28. Works for n = 3, times out for n >= 4. + // (sqrt(x_1) + ... + sqrt(x_n))^2 + // ... = sum_{i <= n} x_i + 2 * sum_{i < j <= n} geo(x_i, x_j) + let build_domain = |n| { + if n < 2 { + panic!("n must be >= 2"); + } + let mut v = Vec::new(); + for i in 0..n { + v.push((format!("x{}", i), domain::pos_dom())); + } + v + }; + let build_term = |n| { + if n < 2 { + panic!("n must be >= 2"); + } + let mut sqrts = Vec::new(); + for i in 0..n { + sqrts.push(format!("(sqrt (var x{}))", i)); + } + let last = sqrts[n-1].clone(); + let before_last = sqrts[n-2].clone(); + let mut t = format!("(add {} {})", before_last, last); + for i in (0..n-2).rev() { + let s = sqrts[i].clone(); + t = format!("(add {} {})", s, t); + } + format!("(neg (pow {} 2))", t) + }; + let domain_pre = build_domain(3).clone(); + let domain = + domain_pre + .iter() + .map(|(s,d)| (s.as_str(), d.clone())) + .collect::>(); + assert_steps_from_string_with_domain(domain, &build_term(3)); } } From 1abc8c93a210c814f94eaacdb1222b3dace524c1 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 6 Dec 2023 14:45:01 -0500 Subject: [PATCH 067/189] chore: increase node limit and adjust time limit --- rewriter/src/extract.rs | 8 +++++--- 1 file changed, 5 insertions(+), 3 deletions(-) diff --git a/rewriter/src/extract.rs b/rewriter/src/extract.rs index fd89bc0d..ab50fc73 100644 --- a/rewriter/src/extract.rs +++ b/rewriter/src/extract.rs @@ -1,4 +1,5 @@ use egg::{*}; +use std::convert::TryInto; use std::fs; use std::time::Duration; use std::collections::HashMap; @@ -165,20 +166,21 @@ pub fn get_steps_from_string(prob_s: &str, domains_vec: Vec<(String, Domain)>, d return None; } } - - for node_limit in [2500, 5000, 10000, 20000] { + + for node_limit in [2500, 5000, 10000, 25000, 50000] { let analysis = Meta { domains : domains.clone() }; let iter_limit = node_limit / 250; + let time_limit = (node_limit / 500).try_into().unwrap(); let runner: Runner = Runner::new(analysis) .with_explanations_enabled() .with_explanation_length_optimization() .with_node_limit(node_limit) .with_iter_limit(iter_limit) - .with_time_limit(Duration::from_secs(5)) + .with_time_limit(Duration::from_secs(time_limit)) .with_expr(&expr) .run(&rules()); From 4dc9d32d534c4e88d02f0a0c08d3707e4c5e8786 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 6 Dec 2023 14:47:52 -0500 Subject: [PATCH 068/189] feat: add useful simplification `binomial_two` --- CvxLean/Tactic/Convexify/RewriteMapLibrary.lean | 6 ++++++ rewriter/src/rules.rs | 2 ++ 2 files changed, 8 insertions(+) diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index 673324f3..a3f4a01a 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -29,6 +29,10 @@ lemma Real.pow_half_two {x : ℝ} (hx : 0 ≤ x) : (x ^ (1 / 2)) ^ 2 = x := by rw [rpow_eq_pow, rpow_eq_pow, ← rpow_mul hx] norm_num +-- TODO: Move. +lemma Real.binomial_two (x y : ℝ) : (x + y) ^ 2 = x ^ 2 + (2 * (x * y) + y ^ 2) := by + simp only [rpow_two]; ring + -- TODO: Move. lemma Real.exp_neg_eq_one_div (x : ℝ) : exp (-x) = 1 / exp x := by rw [exp_neg, inv_eq_one_div] @@ -245,6 +249,8 @@ register_rewrite_map "pow_half_two"; "(pow (pow ?a 0.5) 2)" => "?a" := register_rewrite_map "pow_half_two-rev"; "?a" => "(pow (pow ?a 0.5) 2)" := simp_or_rw [Real.pow_half_two (by positivity_ext)]; +register_rewrite_map "binomial_two"; "(pow (add ?a ?b) 2)" => "(add (pow ?a 2) (add (mul 2 (mul ?a ?b)) (pow ?b 2)))" := + simp_or_rw [Real.binomial_two]; /- Exponential and logarithm rules. -/ diff --git a/rewriter/src/rules.rs b/rewriter/src/rules.rs index 47f69dd7..7098c518 100644 --- a/rewriter/src/rules.rs +++ b/rewriter/src/rules.rs @@ -165,6 +165,8 @@ pub fn rules() -> Vec> { vec![ rw!("pow_half_two"; "(pow (pow ?a 0.5) 2)" => "?a" if is_ge_zero("?a")), rw!("pow_half_two-rev"; "?a" => "(pow (pow ?a 0.5) 2)" if is_ge_zero("?a")), + + rw!("binomial_two"; "(pow (add ?a ?b) 2)" => "(add (pow ?a 2) (add (mul 2 (mul ?a ?b)) (pow ?b 2)))"), /* Exponential and logarithm rules. */ From 67579ee3300614f6a80495a637d5891185451711 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 6 Dec 2023 14:56:52 -0500 Subject: [PATCH 069/189] chore: move Rust tests and create `lib.rs` --- rewriter/src/lib.rs | 14 ++++++++++++++ rewriter/src/main.rs | 4 ---- rewriter/{src => tests}/tests.rs | 23 +++++++++-------------- rewriter/{src => tests}/tests_domain.rs | 7 ++----- 4 files changed, 25 insertions(+), 23 deletions(-) create mode 100644 rewriter/src/lib.rs rename rewriter/{src => tests}/tests.rs (96%) rename rewriter/{src => tests}/tests_domain.rs (99%) diff --git a/rewriter/src/lib.rs b/rewriter/src/lib.rs new file mode 100644 index 00000000..e9f33bd6 --- /dev/null +++ b/rewriter/src/lib.rs @@ -0,0 +1,14 @@ +#[allow(dead_code)] +pub mod domain; + +pub mod curvature; + +pub mod optimization; + +pub mod rules; + +pub mod cost; + +pub mod explain_utils; + +pub mod extract; diff --git a/rewriter/src/main.rs b/rewriter/src/main.rs index f93a574c..95c2a1b2 100644 --- a/rewriter/src/main.rs +++ b/rewriter/src/main.rs @@ -20,10 +20,6 @@ use extract::Minimization as Minimization; use extract::Step as Step; use extract::get_steps as get_steps; -mod tests; - -mod tests_domain; - // NOTE: Taken from https://github.com/opencompl/egg-tactic-code #[derive(Deserialize, Debug)] diff --git a/rewriter/src/tests.rs b/rewriter/tests/tests.rs similarity index 96% rename from rewriter/src/tests.rs rename to rewriter/tests/tests.rs index ed39b681..f4611a95 100644 --- a/rewriter/src/tests.rs +++ b/rewriter/tests/tests.rs @@ -1,10 +1,7 @@ -#[cfg(test)] -mod tests { - -use crate::domain; +use egg_convexify::domain; use domain::Domain as Domain; -use crate::extract; +use egg_convexify::extract; use extract::Minimization as Minimization; use extract::get_steps as get_steps; use extract::get_steps_from_string as get_steps_from_string; @@ -328,13 +325,13 @@ fn test_3_36_a() { "(sqrt (add 1 (add (mul 4 (pow (var x) 2)) (mul 16 (pow (var y) 2)))))"); } -#[test] -fn test_3_36_c() { - // log(e^(2x + 3) + e^(4y + 5)) = lse(2x + 3, 4y + 5) - assert_steps_from_string_with_domain( - vec![("x", domain::free_dom()), ("y", domain::free_dom())], - "(log (add (exp (add (mul 2 (var x)) 3)) (exp (add (mul 4 (var y)) 5))))"); -} +// #[test] +// fn test_3_36_c() { +// // log(e^(2x + 3) + e^(4y + 5)) = lse(2x + 3, 4y + 5) +// assert_steps_from_string_with_domain( +// vec![("x", domain::free_dom()), ("y", domain::free_dom())], +// "(log (add (exp (add (mul 2 (var x)) 3)) (exp (add (mul 4 (var y)) 5))))"); +// } #[test] fn test_3_38_e() { @@ -385,5 +382,3 @@ fn test_3_67() { .collect::>(); assert_steps_from_string_with_domain(domain, &build_term(3)); } - -} diff --git a/rewriter/src/tests_domain.rs b/rewriter/tests/tests_domain.rs similarity index 99% rename from rewriter/src/tests_domain.rs rename to rewriter/tests/tests_domain.rs index 434cba42..27809276 100644 --- a/rewriter/src/tests_domain.rs +++ b/rewriter/tests/tests_domain.rs @@ -1,7 +1,6 @@ -#[cfg(test)] -mod tests_domain { +use egg_convexify; -use crate::domain; +use egg_convexify::domain; use domain::Domain as Domain; @@ -683,5 +682,3 @@ fn exp_neg() { let expected = Domain::make_oo(domain::zero(), domain::one()); assert!(result.eq(&expected)); } - -} From 5a02f107e3ee4dc0fa343e16ffca605bff5bdb68 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 6 Dec 2023 14:57:29 -0500 Subject: [PATCH 070/189] chore: remove print in `pow` --- rewriter/src/domain.rs | 1 - 1 file changed, 1 deletion(-) diff --git a/rewriter/src/domain.rs b/rewriter/src/domain.rs index 3569e44b..9aa78171 100644 --- a/rewriter/src/domain.rs +++ b/rewriter/src/domain.rs @@ -773,7 +773,6 @@ pub fn pow(d1: &Domain, d2: &Domain) -> Domain { } else if n == 1 { return d1.clone(); } else if n % 2 == 0 { - println!("here {}", n); return nonneg_dom(); } else { return free_dom(); From 84a31d7c21f2548ad6ab13507c57d0bd7ccd02ad Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 6 Dec 2023 16:14:55 -0500 Subject: [PATCH 071/189] chore: split Rust tests --- rewriter/src/lib.rs | 2 + rewriter/src/test_util.rs | 78 ++++ rewriter/tests/test_almost_dgp.rs | 11 + rewriter/tests/test_cost_function.rs | 19 + .../tests/{tests_domain.rs => test_domain.rs} | 0 rewriter/tests/test_dqcp.rs | 14 + rewriter/tests/test_gp.rs | 44 ++ rewriter/tests/test_main_example.rs | 13 + rewriter/tests/test_misc.rs | 17 + rewriter/tests/test_rules.rs | 129 ++++++ rewriter/tests/test_stanford.rs | 91 +++++ rewriter/tests/tests.rs | 384 ------------------ 12 files changed, 418 insertions(+), 384 deletions(-) create mode 100644 rewriter/src/test_util.rs create mode 100644 rewriter/tests/test_almost_dgp.rs create mode 100644 rewriter/tests/test_cost_function.rs rename rewriter/tests/{tests_domain.rs => test_domain.rs} (100%) create mode 100644 rewriter/tests/test_dqcp.rs create mode 100644 rewriter/tests/test_gp.rs create mode 100644 rewriter/tests/test_main_example.rs create mode 100644 rewriter/tests/test_misc.rs create mode 100644 rewriter/tests/test_rules.rs create mode 100644 rewriter/tests/test_stanford.rs delete mode 100644 rewriter/tests/tests.rs diff --git a/rewriter/src/lib.rs b/rewriter/src/lib.rs index e9f33bd6..5c7ee70c 100644 --- a/rewriter/src/lib.rs +++ b/rewriter/src/lib.rs @@ -12,3 +12,5 @@ pub mod cost; pub mod explain_utils; pub mod extract; + +pub mod test_util; diff --git a/rewriter/src/test_util.rs b/rewriter/src/test_util.rs new file mode 100644 index 00000000..c4d4f6ec --- /dev/null +++ b/rewriter/src/test_util.rs @@ -0,0 +1,78 @@ +use crate::domain; +use domain::Domain as Domain; + +use crate::extract; +use extract::Minimization as Minimization; +use extract::get_steps as get_steps; +use extract::get_steps_from_string as get_steps_from_string; + +fn make(obj: &str, constrs: Vec<&str>) -> Minimization { + let mut constrs_s = Vec::new(); + for i in 0..constrs.len() { + let tag = format!("h{}", i); + constrs_s.push((tag, constrs[i].to_string())); + } + return Minimization { + obj_fun : obj.to_string(), + constrs : constrs_s, + }; +} + +fn convexify_check_with_domain_maybe_print(domains : Vec<(&str, Domain)>, obj: &str, constrs: Vec<&str>, print: bool) { + let prob = make(obj, constrs); + let domains = + domains.iter().map(|(s, d)| ((*s).to_string(), d.clone())).collect(); + let steps = get_steps(prob, domains, true); + if steps.is_none() { + panic!("Test failed, could not rewrite target into DCP form."); + } + if print { + println!("{:?}", steps); + } +} + +pub fn convexify_check_with_domain(domains : Vec<(&str, Domain)>, obj: &str, constrs: Vec<&str>) { + convexify_check_with_domain_maybe_print(domains, obj, constrs, false) +} + +pub fn convexify_check_with_domain_and_print(domains : Vec<(&str, Domain)>, obj: &str, constrs: Vec<&str>) { + convexify_check_with_domain_maybe_print(domains, obj, constrs, true) +} + +pub fn convexify_check(obj: &str, constrs: Vec<&str>) { + convexify_check_with_domain_maybe_print(vec![], obj, constrs, false) +} + +pub fn convexify_check_and_print(obj: &str, constrs: Vec<&str>) { + convexify_check_with_domain_maybe_print(vec![], obj, constrs, true) +} + +// Used to test out-of-context expressions. + +fn convexify_check_expression_with_domain_maybe_print(domains : Vec<(&str, Domain)>, s: &str, print: bool) { + let domains = + domains.iter().map(|(s, d)| ((*s).to_string(), d.clone())).collect(); + let steps = get_steps_from_string(s, domains, true); + if steps.is_none() { + panic!("Test failed, could not rewrite target into DCP form."); + } + if print { + println!("{:?}", steps); + } +} + +pub fn convexify_check_expression_with_domain(domains : Vec<(&str, Domain)>,s: &str) { + convexify_check_expression_with_domain_maybe_print(domains, s, false); +} + +pub fn convexify_check_expression_with_domain_and_print(domains : Vec<(&str, Domain)>,s: &str) { + convexify_check_expression_with_domain_maybe_print(domains, s, true); +} + +pub fn convexify_check_expression(s: &str) { + convexify_check_expression_with_domain_maybe_print(vec![], s, false); +} + +pub fn convexify_check_expression_and_print(s: &str) { + convexify_check_expression_with_domain_maybe_print(vec![], s, true); +} diff --git a/rewriter/tests/test_almost_dgp.rs b/rewriter/tests/test_almost_dgp.rs new file mode 100644 index 00000000..f559e588 --- /dev/null +++ b/rewriter/tests/test_almost_dgp.rs @@ -0,0 +1,11 @@ +use egg_convexify::test_util::{*}; + +#[test] +fn test_agp2() { + convexify_check( + "(exp (var x))", + vec![ + "(le (mul (exp (var x)) (exp (var y))) (neg (div 2691 500)))" + ]); + +} diff --git a/rewriter/tests/test_cost_function.rs b/rewriter/tests/test_cost_function.rs new file mode 100644 index 00000000..acdefb69 --- /dev/null +++ b/rewriter/tests/test_cost_function.rs @@ -0,0 +1,19 @@ +use egg_convexify::test_util::{*}; + +#[test] +fn test_cost_function_number_of_variable_occurences() { + convexify_check( + "0", + vec![ + "(le (var x) (sub 1 (var x)))" + ]); +} + +#[test] +fn test_cost_function_number_of_variable_occurences_2() { + convexify_check( + "0", + vec![ + "(le (add (mul 2 (var x)) (var x)) 0)" + ]); +} diff --git a/rewriter/tests/tests_domain.rs b/rewriter/tests/test_domain.rs similarity index 100% rename from rewriter/tests/tests_domain.rs rename to rewriter/tests/test_domain.rs diff --git a/rewriter/tests/test_dqcp.rs b/rewriter/tests/test_dqcp.rs new file mode 100644 index 00000000..ecdd4628 --- /dev/null +++ b/rewriter/tests/test_dqcp.rs @@ -0,0 +1,14 @@ +use egg_convexify::domain; + +use egg_convexify::test_util::{*}; + + +#[test] +fn test_dqcp() { + convexify_check_with_domain( + vec![("x", domain::pos_dom())], + "(var x)", + vec![ + "(le (sqrt (div (var x) (add (var x) 1))) 1)" + ]); +} diff --git a/rewriter/tests/test_gp.rs b/rewriter/tests/test_gp.rs new file mode 100644 index 00000000..48e05afa --- /dev/null +++ b/rewriter/tests/test_gp.rs @@ -0,0 +1,44 @@ +use egg_convexify::domain; + +use egg_convexify::test_util::{*}; + +#[test] +fn test_gp3() { + convexify_check( + "(exp (var x))", + vec![ + "(le (sqrt (add (mul (exp (var x)) (exp (var x))) (exp (var y)))) 1)" + ]); +} + +#[test] +fn test_gp4() { + convexify_check( + "(div 1 (div (exp (var x)) (exp (var y))))", + vec![ + "(le 2 (exp (var x)))", + "(le (exp (var x)) 3)", + "(le (add (pow (exp (var x)) 2) (div (mul 3 (exp (var y))) (exp (var z)))) (sqrt (exp (var x))))", + "(eq (div (exp (var x)) (exp (var y))) (pow (exp (var z)) 2))" + ]); +} + +#[test] +fn test_gp6() { + convexify_check_with_domain( + vec![ + ("Aflr", domain::pos_dom()), + ("α" , domain::pos_dom()), + ("β" , domain::pos_dom()), + ("γ" , domain::pos_dom()), + ("δ" , domain::pos_dom())], + "(div 1 (mul (mul (exp (var h)) (exp (var w))) (exp (var d))))", + vec![ + "(le (mul 2 (add (mul (exp (var h)) (exp (var d))) (mul (exp (var w)) (exp (var d))))) (param Awall))", + "(le (mul (exp (var w)) (exp (var d))) (param Aflr))", + "(le (param α) (div (exp (var h)) (exp (var w))))", + "(le (div (exp (var h)) (exp (var w))) (param β))", + "(le (param γ) (div (exp (var d)) (exp (var w))))", + "(le (div (exp (var d)) (exp (var w))) (param δ))" + ]); +} diff --git a/rewriter/tests/test_main_example.rs b/rewriter/tests/test_main_example.rs new file mode 100644 index 00000000..607906a3 --- /dev/null +++ b/rewriter/tests/test_main_example.rs @@ -0,0 +1,13 @@ +use egg_convexify::domain; + +use egg_convexify::test_util::{*}; + +#[test] +fn test_main_example() { + convexify_check_with_domain( + vec![("x", domain::pos_dom())], + "0", + vec![ + "(le (div 1 (sqrt (var x))) (exp (var x)))" + ]); +} diff --git a/rewriter/tests/test_misc.rs b/rewriter/tests/test_misc.rs new file mode 100644 index 00000000..1dc2aaba --- /dev/null +++ b/rewriter/tests/test_misc.rs @@ -0,0 +1,17 @@ +use egg_convexify::domain; + +use egg_convexify::test_util::{*}; + +#[test] +fn test_norm2_with_one() { + convexify_check_expression_with_domain( + vec![("x", domain::free_dom())], + "(sqrt (add (pow (var x) 2) 1))"); +} + +#[test] +fn test_sqrt_pow4() { + convexify_check_expression_with_domain( + vec![("x", domain::nonneg_dom())], + "(sqrt (pow (var x) 4))"); +} diff --git a/rewriter/tests/test_rules.rs b/rewriter/tests/test_rules.rs new file mode 100644 index 00000000..d859e0a6 --- /dev/null +++ b/rewriter/tests/test_rules.rs @@ -0,0 +1,129 @@ +use egg_convexify::domain; +use egg_convexify::domain::Domain as Domain; + +use egg_convexify::test_util::{*}; + +#[test] +fn test_log_le_log() { + convexify_check_with_domain( + vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], + "0", + vec![ + "(le (log (var x)) (log (var y)))" + ]); +} + +#[test] +fn test_sub_iff_add_le() { + convexify_check( + "0", + vec![ + "(le (add 1 (var x)) (var x))", + ]) +} + +#[test] +fn test_log_le_log_rev() { + convexify_check( + "0", + vec![ + "(le (exp (var x)) (exp (var y)))" + ]); +} + +#[test] +fn test_exp_add() { + convexify_check_with_domain( + vec![("x", domain::pos_dom())], + "0", + vec![ + "(le (exp (add (log (var x)) 2)) 1)" + ]); +} + +#[test] +fn test_exp_neg_eq_one_div_obj() { + convexify_check_with_domain( + vec![("x", Domain::make_ci(domain::one()))], + "(mul (var x) (exp (neg (log (var x)))))", + vec![ + ]); +} + +#[test] +fn test_exp_neg_eq_one_div_constr() { + convexify_check_with_domain( + vec![("x", Domain::make_ci(domain::one()))], + "(le (mul (var x) (exp (neg (log (var x))))) (var x))", + vec![ + ]); +} + +#[test] +fn test_log_mul_rev_constr() { + convexify_check_with_domain( + vec![("x", domain::pos_dom())], + "0", + vec![ + "(le (exp (add (log (var x)) (log (add (var x) 1)))) 1)" + ]); +} + +#[test] +fn test_exp_neg_eq_one_div_rev() { + convexify_check( + "(div 1 (exp (var x)))", + vec![ + "(le 1 (var x))" + ]); +} + +#[test] +fn test_div_self() { + convexify_check_with_domain( + vec![("x", domain::pos_dom())], + "0", + vec![ + "(le (mul (div (var x) (var x)) (var y)) 1)" + ]); +} + +#[test] +fn test_div_le_iff_rev() { + convexify_check_with_domain( + vec![("x", domain::pos_dom())], + "0", + vec![ + "(le (mul (var x) (var y)) (var x))" + ]); +} + +#[test] +fn test_log_div_rev_obj() { + convexify_check_with_domain( + vec![("x", domain::pos_dom())], + "(neg (sub (log (pow (var x) 2)) (log (var x))))", + vec![ + ]); +} + +#[test] +fn test_geo_mean_fold() { + convexify_check_expression_with_domain( + vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], + "(neg (sqrt (mul (var x) (var y))))"); +} + +#[test] +fn test_quad_over_lin_fold() { + convexify_check_expression_with_domain( + vec![("x", domain::free_dom()), ("y", domain::pos_dom())], + "(div (pow (var x) 2) (var y))"); +} + +#[test] +fn test_norm2_fold() { + convexify_check_expression_with_domain( + vec![("x", domain::free_dom()), ("y", domain::free_dom())], + "(sqrt (add (pow (var x) 2) (pow (var y) 2)))"); +} diff --git a/rewriter/tests/test_stanford.rs b/rewriter/tests/test_stanford.rs new file mode 100644 index 00000000..48a974f1 --- /dev/null +++ b/rewriter/tests/test_stanford.rs @@ -0,0 +1,91 @@ +use egg_convexify::domain; + +use egg_convexify::test_util::{*}; + +#[test] +fn test_3_32() { + // 1 / (x * y) = (x^-0.5)^2 / y + // = qol(x^-0.5, y) (this is just one possibility) + convexify_check_expression_with_domain( + vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], + "(div 1 (mul (var x) (var y)))"); +} + +#[test] +fn test_3_33() { + // sqrt(1 + x^4 / y) = sqrt(1^2 + (x^2 / y)^2) + // ... = norm2(1, qol(x, y)) + // NOTE(RFM): Constant unfolding issue: x cannot be in the free domain as + // otherwise we cannot convert x^4 to (x^2)^2. + convexify_check_expression_with_domain( + vec![("x", domain::nonneg_dom()), ("y", domain::pos_dom())], + "(sqrt (add 1 (div (pow (var x) 4) (var y))))" + ); +} + +#[test] +fn test_3_36_a() { + // sqrt(1 + 4x^2 + 16y^2) = sqrt((2x)^2 + (sqrt(1^2 + (4y)^2))^2) + // ... = norm2(2x, norm2(1, 4y)) + convexify_check_expression_with_domain( + vec![("x", domain::nonneg_dom()), ("y", domain::nonneg_dom())], + "(sqrt (add 1 (add (mul 4 (pow (var x) 2)) (mul 16 (pow (var y) 2)))))"); +} + +// #[test] +// fn test_3_36_c() { +// // log(e^(2x + 3) + e^(4y + 5)) = lse(2x + 3, 4y + 5) +// convexify_check_expression_with_domain( +// vec![("x", domain::free_dom()), ("y", domain::free_dom())], +// "(log (add (exp (add (mul 2 (var x)) 3)) (exp (add (mul 4 (var y)) 5))))"); +// } + +#[test] +fn test_3_38_e() { + // (sqrt(x) + sqrt(y))^2 = x + y + 2sqrt(xy) + // ... = x + y + 2geo(x, y) + convexify_check_expression_with_domain( + vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], + "(neg (pow (add (sqrt (var x)) (sqrt (var y))) 2))"); +} + +#[test] +fn test_3_67() { + // Generalizaiton of 3.28. Works for n = 3, times out for n >= 4. + // (sqrt(x_1) + ... + sqrt(x_n))^2 + // ... = sum_{i <= n} x_i + 2 * sum_{i < j <= n} geo(x_i, x_j) + let build_domain = |n| { + if n < 2 { + panic!("n must be >= 2"); + } + let mut v = Vec::new(); + for i in 0..n { + v.push((format!("x{}", i), domain::pos_dom())); + } + v + }; + let build_term = |n| { + if n < 2 { + panic!("n must be >= 2"); + } + let mut sqrts = Vec::new(); + for i in 0..n { + sqrts.push(format!("(sqrt (var x{}))", i)); + } + let last = sqrts[n-1].clone(); + let before_last = sqrts[n-2].clone(); + let mut t = format!("(add {} {})", before_last, last); + for i in (0..n-2).rev() { + let s = sqrts[i].clone(); + t = format!("(add {} {})", s, t); + } + format!("(neg (pow {} 2))", t) + }; + let domain_pre = build_domain(3).clone(); + let domain = + domain_pre + .iter() + .map(|(s,d)| (s.as_str(), d.clone())) + .collect::>(); + convexify_check_expression_with_domain(domain, &build_term(3)); +} diff --git a/rewriter/tests/tests.rs b/rewriter/tests/tests.rs deleted file mode 100644 index f4611a95..00000000 --- a/rewriter/tests/tests.rs +++ /dev/null @@ -1,384 +0,0 @@ -use egg_convexify::domain; -use domain::Domain as Domain; - -use egg_convexify::extract; -use extract::Minimization as Minimization; -use extract::get_steps as get_steps; -use extract::get_steps_from_string as get_steps_from_string; - -fn make(obj: &str, constrs: Vec<&str>) -> Minimization { - let mut constrs_s = Vec::new(); - for i in 0..constrs.len() { - let tag = format!("h{}", i); - constrs_s.push((tag, constrs[i].to_string())); - } - return Minimization { - obj_fun : obj.to_string(), - constrs : constrs_s, - }; -} - -fn assert_steps_with_domain(domains : Vec<(&str, Domain)>, obj: &str, constrs: Vec<&str>) { - let prob = make(obj, constrs); - let domains = - domains.iter().map(|(s, d)| ((*s).to_string(), d.clone())).collect(); - let steps = get_steps(prob, domains, true); - println!("{:?}", steps); - assert!(steps.is_some()); -} - -fn assert_steps(obj: &str, constrs: Vec<&str>) { - assert_steps_with_domain(vec![], obj, constrs); -} - -// Used to test single expressions outside the context of an optimization -// problem. -fn assert_steps_from_string_with_domain(domains : Vec<(&str, Domain)>, s: &str) { - let domains = - domains.iter().map(|(s, d)| ((*s).to_string(), d.clone())).collect(); - let steps = get_steps_from_string(s, domains, true); - println!("{:?}", steps); - assert!(steps.is_some()); -} - -#[allow(unused)] -fn assert_steps_from_string(s: &str) { - assert_steps_from_string_with_domain(vec![], s); -} - - -// Examples. - -#[test] -fn test_main_example() { - assert_steps_with_domain( - vec![("x", domain::pos_dom())], - "0", - vec![ - "(le (div 1 (sqrt (var x))) (exp (var x)))" - ]); - -} - -#[test] -fn test_agp2() { - assert_steps( - "(exp (var x))", - vec![ - "(le (mul (exp (var x)) (exp (var y))) (neg (div 2691 500)))" - ]); - -} - -#[test] -fn test_gp3() { - assert_steps( - "(exp (var x))", - vec![ - "(le (sqrt (add (mul (exp (var x)) (exp (var x))) (exp (var y)))) 1)" - ]); -} - -#[test] -fn test_gp4() { - assert_steps( - "(div 1 (div (exp (var x)) (exp (var y))))", - vec![ - "(le 2 (exp (var x)))", - "(le (exp (var x)) 3)", - "(le (add (pow (exp (var x)) 2) (div (mul 3 (exp (var y))) (exp (var z)))) (sqrt (exp (var x))))", - "(eq (div (exp (var x)) (exp (var y))) (pow (exp (var z)) 2))" - ]); -} - -#[test] -fn test_gp6() { - assert_steps_with_domain( - vec![("Aflr", domain::pos_dom()), ("α", domain::pos_dom()), ("β", domain::pos_dom()), ("γ", domain::pos_dom()), ("δ", domain::pos_dom())], - "(div 1 (mul (mul (exp (var h)) (exp (var w))) (exp (var d))))", - vec![ - "(le (mul 2 (add (mul (exp (var h)) (exp (var d))) (mul (exp (var w)) (exp (var d))))) (param Awall))", - "(le (mul (exp (var w)) (exp (var d))) (param Aflr))", - "(le (param α) (div (exp (var h)) (exp (var w))))", - "(le (div (exp (var h)) (exp (var w))) (param β))", - "(le (param γ) (div (exp (var d)) (exp (var w))))", - "(le (div (exp (var d)) (exp (var w))) (param δ))" - ]); -} - - -#[test] -fn test_dqcp() { - assert_steps_with_domain( - vec![("x", domain::pos_dom())], - "(var x)", - vec![ - "(le (sqrt (div (var x) (add (var x) 1))) 1)" - ]); -} - -// Cost function. - -#[test] -fn test_cost_function_number_of_variable_occurences() { - assert_steps( - "0", - vec![ - "(le (var x) (sub 1 (var x)))" - ]); -} - -#[test] -fn test_cost_function_number_of_variable_occurences_2() { - assert_steps( - "0", - vec![ - "(le (add (mul 2 (var x)) (var x)) 0)" - ]); -} - -#[test] -fn test_cost_function_number_of_variable_occurences_3() { - assert_steps( - "0", - vec![ - "(le (add (mul 2 (var x)) (mul 3 (var x))) 0)" - ]); -} - -// Rule-based tests. - -#[test] -fn test_log_le_log() { - assert_steps_with_domain( - vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], - "0", - vec![ - "(le (log (var x)) (log (var y)))" - ]); -} - -#[test] -fn test_sub_iff_add_le() { - assert_steps( - "0", - vec![ - "(le (add 1 (var x)) (var x))", - ]) -} - -#[test] -fn test_log_le_log_rev() { - assert_steps( - "0", - vec![ - "(le (exp (var x)) (exp (var y)))" - ]); -} - -#[test] -fn test_exp_add() { - assert_steps_with_domain( - vec![("x", domain::pos_dom())], - "0", - vec![ - "(le (exp (add (log (var x)) 2)) 1)" - ]); -} - -#[test] -fn test_exp_neg_eq_one_div_obj() { - assert_steps_with_domain( - vec![("x", Domain::make_ci(domain::one()))], - "(mul (var x) (exp (neg (log (var x)))))", - vec![ - ]); -} - -#[test] -fn test_exp_neg_eq_one_div_constr() { - assert_steps_with_domain( - vec![("x", Domain::make_ci(domain::one()))], - "(le (mul (var x) (exp (neg (log (var x))))) (var x))", - vec![ - ]); -} - -#[test] -fn test_log_mul_rev_constr() { - assert_steps_with_domain( - vec![("x", domain::pos_dom())], - "0", - vec![ - "(le (exp (add (log (var x)) (log (add (var x) 1)))) 1)" - ]); -} - -#[test] -fn test_exp_neg_eq_one_div_rev() { - assert_steps( - "(div 1 (exp (var x)))", - vec![ - "(le 1 (var x))" - ]); -} - -#[test] -fn test_div_self() { - assert_steps_with_domain( - vec![("x", domain::pos_dom())], - "0", - vec![ - "(le (mul (div (var x) (var x)) (var y)) 1)" - ]); -} - -#[test] -fn test_div_le_iff_rev() { - assert_steps_with_domain( - vec![("x", domain::pos_dom())], - "0", - vec![ - "(le (mul (var x) (var y)) (var x))" - ]); -} - -#[test] -fn test_log_div_rev_obj() { - assert_steps_with_domain( - vec![("x", domain::pos_dom())], - "(neg (sub (log (pow (var x) 2)) (log (var x))))", - vec![ - ]); -} - -// Folding atoms. - -#[test] -fn test_geo_mean() { - assert_steps_from_string_with_domain( - vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], - "(neg (sqrt (mul (var x) (var y))))"); -} - -#[test] -fn test_quad_over_lin() { - assert_steps_from_string_with_domain( - vec![("x", domain::free_dom()), ("y", domain::pos_dom())], - "(div (pow (var x) 2) (var y))"); -} - -#[test] -fn test_norm2() { - assert_steps_from_string_with_domain( - vec![("x", domain::free_dom()), ("y", domain::free_dom())], - "(sqrt (add (pow (var x) 2) (pow (var y) 2)))"); -} - -// Misc. - -#[test] -fn test_norm2_with_one() { - assert_steps_from_string_with_domain( - vec![("x", domain::free_dom())], - "(sqrt (add (pow (var x) 2) 1))"); -} - -#[test] -fn test_sqrt_pow4() { - assert_steps_from_string_with_domain( - vec![("x", domain::nonneg_dom())], - "(sqrt (pow (var x) 4))"); -} - -// From Stanford course, additional exercises. - -#[test] -fn test_3_32() { - // 1 / (x * y) = (x^-0.5)^2 / y - // = qol(x^-0.5, y) (this is just one possibility) - assert_steps_from_string_with_domain( - vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], - "(div 1 (mul (var x) (var y)))"); -} - - - -#[test] -fn test_3_33() { - // sqrt(1 + x^4 / y) = sqrt(1^2 + (x^2 / y)^2) - // ... = norm2(1, qol(x, y)) - // NOTE(RFM): Constant unfolding issue: x cannot be in the free domain as - // otherwise we cannot convert x^4 to (x^2)^2. - assert_steps_from_string_with_domain( - vec![("x", domain::nonneg_dom()), ("y", domain::pos_dom())], - "(sqrt (add 1 (div (pow (var x) 4) (var y))))" - ); -} - -#[test] -fn test_3_36_a() { - // sqrt(1 + 4x^2 + 16y^2) = sqrt((2x)^2 + (sqrt(1^2 + (4y)^2))^2) - // ... = norm2(2x, norm2(1, 4y)) - assert_steps_from_string_with_domain( - vec![("x", domain::nonneg_dom()), ("y", domain::nonneg_dom())], - "(sqrt (add 1 (add (mul 4 (pow (var x) 2)) (mul 16 (pow (var y) 2)))))"); -} - -// #[test] -// fn test_3_36_c() { -// // log(e^(2x + 3) + e^(4y + 5)) = lse(2x + 3, 4y + 5) -// assert_steps_from_string_with_domain( -// vec![("x", domain::free_dom()), ("y", domain::free_dom())], -// "(log (add (exp (add (mul 2 (var x)) 3)) (exp (add (mul 4 (var y)) 5))))"); -// } - -#[test] -fn test_3_38_e() { - // (sqrt(x) + sqrt(y))^2 = x + y + 2sqrt(xy) - // ... = x + y + 2geo(x, y) - assert_steps_from_string_with_domain( - vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], - "(neg (pow (add (sqrt (var x)) (sqrt (var y))) 2))"); -} - -#[test] -fn test_3_67() { - // Generalizaiton of 3.28. Works for n = 3, times out for n >= 4. - // (sqrt(x_1) + ... + sqrt(x_n))^2 - // ... = sum_{i <= n} x_i + 2 * sum_{i < j <= n} geo(x_i, x_j) - let build_domain = |n| { - if n < 2 { - panic!("n must be >= 2"); - } - let mut v = Vec::new(); - for i in 0..n { - v.push((format!("x{}", i), domain::pos_dom())); - } - v - }; - let build_term = |n| { - if n < 2 { - panic!("n must be >= 2"); - } - let mut sqrts = Vec::new(); - for i in 0..n { - sqrts.push(format!("(sqrt (var x{}))", i)); - } - let last = sqrts[n-1].clone(); - let before_last = sqrts[n-2].clone(); - let mut t = format!("(add {} {})", before_last, last); - for i in (0..n-2).rev() { - let s = sqrts[i].clone(); - t = format!("(add {} {})", s, t); - } - format!("(neg (pow {} 2))", t) - }; - let domain_pre = build_domain(3).clone(); - let domain = - domain_pre - .iter() - .map(|(s,d)| (s.as_str(), d.clone())) - .collect::>(); - assert_steps_from_string_with_domain(domain, &build_term(3)); -} From 8c882a634d4a3026e783382d5ffc30d29a7c6ca5 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 6 Dec 2023 16:18:13 -0500 Subject: [PATCH 072/189] refactor: `explain_utils` -> `explain_util` --- rewriter/src/{explain_utils.rs => explain_util.rs} | 0 rewriter/src/extract.rs | 6 +++--- rewriter/src/lib.rs | 2 +- rewriter/src/main.rs | 2 +- 4 files changed, 5 insertions(+), 5 deletions(-) rename rewriter/src/{explain_utils.rs => explain_util.rs} (100%) diff --git a/rewriter/src/explain_utils.rs b/rewriter/src/explain_util.rs similarity index 100% rename from rewriter/src/explain_utils.rs rename to rewriter/src/explain_util.rs diff --git a/rewriter/src/extract.rs b/rewriter/src/extract.rs index ab50fc73..78edd283 100644 --- a/rewriter/src/extract.rs +++ b/rewriter/src/extract.rs @@ -21,9 +21,9 @@ use rules::rules as rules; use crate::cost; use cost::DCPCost as DCPCost; -use crate::explain_utils; -use explain_utils::Direction as Direction; -use explain_utils::get_rewrite_name_and_direction as get_rewrite_name_and_direction; +use crate::explain_util; +use explain_util::Direction as Direction; +use explain_util::get_rewrite_name_and_direction as get_rewrite_name_and_direction; #[derive(Serialize, Debug)] pub struct Step { diff --git a/rewriter/src/lib.rs b/rewriter/src/lib.rs index 5c7ee70c..cc060f82 100644 --- a/rewriter/src/lib.rs +++ b/rewriter/src/lib.rs @@ -9,7 +9,7 @@ pub mod rules; pub mod cost; -pub mod explain_utils; +pub mod explain_util; pub mod extract; diff --git a/rewriter/src/main.rs b/rewriter/src/main.rs index 95c2a1b2..7638fb77 100644 --- a/rewriter/src/main.rs +++ b/rewriter/src/main.rs @@ -13,7 +13,7 @@ mod rules; mod cost; -mod explain_utils; +mod explain_util; mod extract; use extract::Minimization as Minimization; From d463e8ff7bf003cdaa04d45dbc4bc0af2ab2b17a Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 6 Dec 2023 16:24:55 -0500 Subject: [PATCH 073/189] chore: proper names for solver input and output files --- .gitignore | 5 ++--- CvxLean/Tactic/Solver/Conic.lean | 8 ++++---- 2 files changed, 6 insertions(+), 7 deletions(-) diff --git a/.gitignore b/.gitignore index 51cb7779..979e6119 100644 --- a/.gitignore +++ b/.gitignore @@ -1,9 +1,8 @@ .DS_Store /build -/solver/build -/solver/test.cbf -/solver/test.sol +/solver/*.cbf +/solver/*.sol /rewriter/*.dot /rewriter/*.svg /ffi/build diff --git a/CvxLean/Tactic/Solver/Conic.lean b/CvxLean/Tactic/Solver/Conic.lean index 2ef73ba4..cd1ae907 100644 --- a/CvxLean/Tactic/Solver/Conic.lean +++ b/CvxLean/Tactic/Solver/Conic.lean @@ -92,7 +92,7 @@ unsafe def conicSolverFromValues (goalExprs : SolutionExpr) (data : ProblemData) (CBF.ConeProduct.mk n fixedCones.length fixedCones) -- Write input. - let inputPath := "solver/test.cbf" + let inputPath := "solver/input.cbf" IO.FS.writeFile inputPath (ToString.toString cbf) -- Adjust path to MOSEK. @@ -105,9 +105,10 @@ unsafe def conicSolverFromValues (goalExprs : SolutionExpr) (data : ProblemData) mosekBinPath -- Run solver. + let outputPath := "solver/output.sol" let out ← IO.Process.output { cmd := "mosek", - args := #[inputPath], + args := #[inputPath, "-out", outputPath], env := #[("PATH", p)] } if out.exitCode != 0 then dbg_trace ("MOSEK exited with code " ++ ToString.toString out.exitCode) @@ -117,7 +118,6 @@ unsafe def conicSolverFromValues (goalExprs : SolutionExpr) (data : ProblemData) IO.println res -- Read output. - let outputPath := "solver/test.sol" let handle ← IO.FS.Handle.mk outputPath IO.FS.Mode.read let output ← IO.FS.Handle.readToEnd handle @@ -139,7 +139,7 @@ unsafe def exprFromSol (goalExprs : SolutionExpr) (sol : Sol.Result) : MetaM Exp -- Vectors and matrices as functions. let mut solPointExprArray : Array Expr := #[] - -- TODO: This won't work in general, need to take into account the + -- TODO(RFM): This won't work in general, need to take into account the -- associativity of the variables if there are products. Infer dimension might -- need to return a tree. let mut i : ℕ := 0 From 0ef86ec298cde034895ae8581af37b155e653979 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 6 Dec 2023 16:37:27 -0500 Subject: [PATCH 074/189] feat: support for unrotated second-order cone --- CvxLean/Tactic/Solver/Float/Coeffs.lean | 71 +++++++++-------- CvxLean/Tactic/Solver/Float/ProblemData.lean | 81 +++++++++++--------- 2 files changed, 82 insertions(+), 70 deletions(-) diff --git a/CvxLean/Tactic/Solver/Float/Coeffs.lean b/CvxLean/Tactic/Solver/Float/Coeffs.lean index a9f0ec6b..c801b073 100644 --- a/CvxLean/Tactic/Solver/Float/Coeffs.lean +++ b/CvxLean/Tactic/Solver/Float/Coeffs.lean @@ -35,13 +35,13 @@ unsafe def evalFloat (e : Expr) : MetaM Float := do /-- Generate an array of elements of a finite type -/ unsafe def elemsOfFintype (ty : Expr) : MetaM (Array Expr) := do match ty with - | Expr.app (Expr.const ``Fin _) nExpr => do + | .app (.const ``Fin _) nExpr => do let n : Nat ← evalExpr Nat (mkConst ``Nat) nExpr let mut res := #[] for i in [:n] do res := res.push (← mkFinIdxExpr i n) return res - | Expr.app (Expr.app (Expr.const ``Sum lvl) tyl) tyr => do + | .app (.app (.const ``Sum lvl) tyl) tyr => do let elemsl := (← elemsOfFintype tyl).map fun e => mkAppN (mkConst ``Sum.inl lvl) #[tyl, tyr, e] let elemsr := (← elemsOfFintype tyr).map fun e => @@ -52,13 +52,10 @@ unsafe def elemsOfFintype (ty : Expr) : MetaM (Array Expr) := do /- Evaluate floating point matrix expressions. -/ unsafe def evalFloatMatrix (e : Expr) : MetaM (Array (Array Float)) := do let (tyn, tym) ← do (match (← inferType e) with - | Expr.forallE _ tyn - (Expr.forallE _ tym - (Expr.const ``Float _) _) _ => + | .forallE _ tyn (.forallE _ tym (.const ``Float _) _) _ => return (tyn, tym) - | Expr.app (Expr.app (Expr.app (Expr.const ``Matrix _) - tyn) tym) - (Expr.const ``Float _) => + | .app (.app (.app (.const ``Matrix _) tyn) tym) + (.const ``Float _) => return (tyn, tym) | _ => throwError "Not a float matrix: {e} {e.ctorName}.") let elemsn ← elemsOfFintype tyn @@ -78,18 +75,18 @@ shape. Used to evaluate the constant term in a constraint or objective function. partial def generateZerosOfShape (ty : Expr) : MetaM Expr := match ty.consumeMData with -- 1-dimensional variables. - | Expr.const ``Float _ => + | .const ``Float _ => return (mkFloat 0) -- Vectors. - | Expr.forallE _ ty (Expr.const ``Float _) _ => + | .forallE _ ty (.const ``Float _) _ => return (mkLambda `_ Lean.BinderInfo.default ty (mkFloat 0)) -- Matrices. - | Expr.app (Expr.app (Expr.app (Expr.const ``Matrix _) tyn) tym) - (Expr.const ``Float _) => do + | .app (.app (.app (.const ``Matrix _) tyn) tym) + (.const ``Float _) => do return (mkLambda `_ Lean.BinderInfo.default tyn ((mkLambda `_ Lean.BinderInfo.default tym) (mkFloat 0))) -- Products. - | Expr.app (Expr.app (Expr.const ``Prod _) tyl) tyr => do + | .app (.app (.const ``Prod _) tyl) tyr => do let l ← generateZerosOfShape tyl let r ← generateZerosOfShape tyr return ← mkAppM ``Prod.mk #[l, r] @@ -103,12 +100,12 @@ unsafe def generateBasisOfShape (ty : Expr) : MetaM (Array Expr × Array Expr) := match ty.consumeMData with -- 1-dimensional variables. - | Expr.const ``Float _ => + | .const ``Float _ => return (#[], #[mkFloat 1]) -- Vectors. - | Expr.forallE _ + | .forallE _ tyn - (Expr.const ``Float _) _ => do + (.const ``Float _) _ => do let mut res := #[] for i in ← elemsOfFintype tyn do let b ← withLocalDeclD `i' tyn fun i' => do @@ -118,8 +115,8 @@ unsafe def generateBasisOfShape (ty : Expr) res := res.push b return (#[], res) -- Matrices. - | Expr.app (Expr.app (Expr.app (Expr.const ``Matrix _) tyn) tym) - (Expr.const ``Float _) => do + | .app (.app (.app (.const ``Matrix _) tyn) tym) + (.const ``Float _) => do let mut res := #[] let elemsn ← elemsOfFintype tyn for i in elemsn do @@ -134,7 +131,7 @@ unsafe def generateBasisOfShape (ty : Expr) -- TODO: For now we're treating matrices as a bunch of scalars. return (#[], res) -- Products. - | Expr.app (Expr.app (Expr.const ``Prod _) tyl) tyr => do + | .app (.app (.const ``Prod _) tyl) tyr => do let r₀ ← generateZerosOfShape tyr let l₀ ← generateZerosOfShape tyl @@ -159,22 +156,21 @@ unsafe def unrollVectors (constraints : Expr) : MetaM (Array Expr) := do let c' := Expr.consumeMData c match c' with -- Vector zero cone. - | Expr.app (Expr.app (Expr.const ``Real.Vec.zeroCone _) n) e => + | .app (.app (.const ``Real.Vec.zeroCone _) n) e => let n : Nat ← evalExpr Nat (mkConst ``Nat) n for i in [:n] do let idxExpr ← mkFinIdxExpr i n let ei := mkApp e idxExpr res := res.push (← mkAppM ``Real.zeroCone #[ei]) -- Positive orthant cone. - | Expr.app (Expr.app (Expr.const ``Real.Vec.posOrthCone _) n) e => + | .app (.app (.const ``Real.Vec.posOrthCone _) n) e => let n : Nat ← evalExpr Nat (mkConst ``Nat) n for i in [:n] do let idxExpr ← mkFinIdxExpr i n let ei := mkApp e idxExpr res := res.push (← mkAppM ``Real.posOrthCone #[ei]) -- Vector exponential cone. - | Expr.app (Expr.app (Expr.app (Expr.app - (Expr.const ``Real.Vec.expCone _) n) a) b) c => + | .app (.app (.app (.app (.const ``Real.Vec.expCone _) n) a) b) c => let n : Nat ← evalExpr Nat (mkConst ``Nat) n for i in [:n] do let idxExpr ← mkFinIdxExpr i n @@ -249,15 +245,15 @@ unsafe def determineCoeffsFromExpr (goalExprs : Meta.SolutionExpr) for c in cs do trace[Meta.debug] "Coeffs going through constraint {c}." match Expr.consumeMData c with - | Expr.app (Expr.const ``Real.zeroCone _) e => do + | .app (.const ``Real.zeroCone _) e => do let e ← realToFloat e let res ← determineScalarCoeffsAux e p floatDomain data := data.addZeroConstraint res.1 res.2 - | Expr.app (Expr.const ``Real.posOrthCone _) e => do + | .app (.const ``Real.posOrthCone _) e => do let e ← realToFloat e let res ← determineScalarCoeffsAux e p floatDomain data := data.addPosOrthConstraint res.1 res.2 - | Expr.app (Expr.app (Expr.app (Expr.const ``Real.expCone _) a) b) c => do + | .app (.app (.app (.const ``Real.expCone _) a) b) c => do let res ← #[a, b, c].mapM fun e => do let e ← realToFloat e return ← determineScalarCoeffsAux e p floatDomain @@ -266,9 +262,8 @@ unsafe def determineCoeffsFromExpr (goalExprs : Meta.SolutionExpr) data := data.addExpConstraint res[2]!.1 res[2]!.2 data := data.addExpConstraint res[1]!.1 res[1]!.2 data := data.addExpConstraint res[0]!.1 res[0]!.2 - | Expr.app (Expr.app (Expr.app (Expr.app (Expr.app - (Expr.const ``Real.rotatedSoCone _) - (Expr.app (Expr.const ``Fin _) n)) _) v) w) x => do + | .app (.app (.app (.app (.app (.const ``Real.rotatedSoCone _) + (.app (.const ``Fin _) n)) _) v) w) x => do let n : Nat ← evalExpr Nat (mkConst ``Nat) n -- TODO: This is a common issue with all vectors. let xis ← (Array.range n).mapM @@ -277,9 +272,19 @@ unsafe def determineCoeffsFromExpr (goalExprs : Meta.SolutionExpr) let e ← realToFloat e let (ea, eb) ← determineScalarCoeffsAux e p floatDomain data := data.addRotatedSOConstraint ea eb + | .app (.app (.app (.app + (.const ``Real.soCone _) + (.app (.const ``Fin _) n)) _) t) x => do + let n : Nat ← evalExpr Nat (mkConst ``Nat) n + -- TODO: This is a common issue with all vectors. + let xis ← (Array.range n).mapM + (fun i => return (mkApp x (← mkFinIdxExpr i n))) + for e in (#[t] ++ xis) do + let e ← realToFloat e + let (ea, eb) ← determineScalarCoeffsAux e p floatDomain + data := data.addSOConstraint ea eb -- TODO: Unroll? - | Expr.app (Expr.app (Expr.app - (Expr.const ``Real.Matrix.posOrthCone _) m) n) e => do + | .app (.app (.app (.const ``Real.Matrix.posOrthCone _) m) n) e => do let m : Nat ← evalExpr Nat (mkConst ``Nat) m let n : Nat ← evalExpr Nat (mkConst ``Nat) n for i in [:m] do @@ -289,8 +294,8 @@ unsafe def determineCoeffsFromExpr (goalExprs : Meta.SolutionExpr) let eij ← realToFloat eij let res ← determineScalarCoeffsAux eij p floatDomain data := data.addPosOrthConstraint res.1 res.2 - | Expr.app (Expr.app (Expr.app - (Expr.const ``Real.Matrix.PSDCone _) mty) _) e => do + | .app (.app (.app + (.const ``Real.Matrix.PSDCone _) mty) _) e => do let e ← realToFloat e let res ← determineMatrixCoeffsAux e p floatDomain -- The size of the matrix can be inferred from number of coefficients. diff --git a/CvxLean/Tactic/Solver/Float/ProblemData.lean b/CvxLean/Tactic/Solver/Float/ProblemData.lean index 0a17d331..1ff0d7bb 100644 --- a/CvxLean/Tactic/Solver/Float/ProblemData.lean +++ b/CvxLean/Tactic/Solver/Float/ProblemData.lean @@ -1,23 +1,24 @@ namespace CvxLean -/-- Cones admitting scalar affine constraints. Note that the only cone that +/-- Cones admitting scalar affine constraints. Note that the only cone that admits matrix affine constraints is the PSD cone. -/ -inductive ScalarConeType - | Zero | PosOrth | Exp | QR +inductive ScalarConeType + | Zero | PosOrth | Exp | Q | QR -namespace ScalarConeType +namespace ScalarConeType -instance : ToString ScalarConeType where - toString +instance : ToString ScalarConeType where + toString | Zero => "Zero" | PosOrth => "PosOrth" | Exp => "Exp" + | Q => "Q" | QR => "QR" -end ScalarConeType +end ScalarConeType /-- Encoding the constraint ⟨X, A⟩ + ∑ i, aᵢ xᵢ + b. -/ -structure ScalarAffine := +structure ScalarAffine := (n m : Nat) (A : Array (Array Float)) (a : Array Float) @@ -25,67 +26,67 @@ structure ScalarAffine := namespace ScalarAffine -instance : ToString ScalarAffine where - toString sa := +instance : ToString ScalarAffine where + toString sa := s!"ScalarAffine [{sa.n}, {sa.m}, {sa.A}, {sa.a}, {sa.b}]" end ScalarAffine /-- Encoding the constraint ∑ i, xᵢ • Hᵢ + D -/ -structure MatrixAffine := +structure MatrixAffine := (n : Nat) (H : Array (Array (Array Float))) (D : Array (Array Float)) namespace MatrixAffine -instance : ToString MatrixAffine where - toString sa := +instance : ToString MatrixAffine where + toString sa := s!"MatrixAffine [{sa.n}, {sa.H}, {sa.D}]" end MatrixAffine /-- Coeffs generates this from the problem goal. -/ -structure ProblemData := +structure ProblemData := (objective : Option ScalarAffine) (scalarAffineConstraints : Array (ScalarAffine × ScalarConeType)) (matrixAffineConstraints : Array MatrixAffine) namespace ProblemData -instance : ToString ProblemData where - toString data := - let scalarConstraintsStr := data.scalarAffineConstraints.map +instance : ToString ProblemData where + toString data := + let scalarConstraintsStr := data.scalarAffineConstraints.map (fun (sa, sct) => (toString sct) ++ " " ++ (toString sa)) - let matrixConstraintsStr := data.matrixAffineConstraints.map + let matrixConstraintsStr := data.matrixAffineConstraints.map (fun ma => "PSD " ++ (toString ma)) let constraintsStr := (scalarConstraintsStr ++ matrixConstraintsStr).map (fun s => "* " ++ s ++ "\n") - "Objective: " ++ toString (data.objective) ++ "\n" ++ - "Constraints: \n" ++ String.join constraintsStr.data + "Objective: " ++ toString (data.objective) ++ "\n" ++ + "Constraints: \n" ++ String.join constraintsStr.data -def empty : ProblemData := +def empty : ProblemData := ProblemData.mk none #[] #[] -instance : Inhabited ProblemData where +instance : Inhabited ProblemData where default := empty /-- Set full objective function. -/ def setObjective (data : ProblemData) (A : Array (Array Float)) (a : Array Float) (b : Float) - : ProblemData := + : ProblemData := { data with objective := ScalarAffine.mk A.size a.size A a b } /-- Same but only ∑ i, aᵢxᵢ + b. -/ def setObjectiveOnlyVector (data : ProblemData) (a : Array Float) (b : Float) - : ProblemData := + : ProblemData := data.setObjective #[] a b /-- Same but only ⟨A, X⟩ + b. -/ def setObjectiveOnlyMatrix (data : ProblemData) (A : Array (Array Float)) (b : Float) - : ProblemData := + : ProblemData := data.setObjective A #[] b /-- Add a full scalar affine constraint to the problem data. -/ @@ -93,7 +94,7 @@ def addScalarAffineConstraint (data : ProblemData) (A : Array (Array Float)) (a : Array Float) (b : Float) (sct : ScalarConeType) : ProblemData := let constraint := ScalarAffine.mk A.size a.size A a b - { data with scalarAffineConstraints := + { data with scalarAffineConstraints := data.scalarAffineConstraints.push ⟨constraint, sct⟩ } /-- Add a scalar affine constraint without the ⟨A, X⟩ part. -/ @@ -104,36 +105,42 @@ def addScalarAffineConstraintOnlyVector (data : ProblemData) /-- Specialized to the zero cone. -/ def addZeroConstraint (data : ProblemData) - (a : Array Float) (b : Float) - : ProblemData := + (a : Array Float) (b : Float) + : ProblemData := data.addScalarAffineConstraintOnlyVector a b ScalarConeType.Zero /-- Specialized to the exponential cone. -/ def addExpConstraint (data : ProblemData) - (a : Array Float) (b : Float) - : ProblemData := + (a : Array Float) (b : Float) + : ProblemData := data.addScalarAffineConstraintOnlyVector a b ScalarConeType.Exp /-- Specialized to the positive orthant cone. -/ def addPosOrthConstraint (data : ProblemData) - (a : Array Float) (b : Float) - : ProblemData := + (a : Array Float) (b : Float) + : ProblemData := data.addScalarAffineConstraintOnlyVector a b ScalarConeType.PosOrth +/- Specialized to the second-order cone. -/ +def addSOConstraint (data : ProblemData) + (a : Array Float) (b : Float) + : ProblemData := + data.addScalarAffineConstraintOnlyVector a b ScalarConeType.Q + /- Specialized to the quadratic rotated cone. -/ def addRotatedSOConstraint (data : ProblemData) - (a : Array Float) (b : Float) - : ProblemData := + (a : Array Float) (b : Float) + : ProblemData := data.addScalarAffineConstraintOnlyVector a b ScalarConeType.QR /-- Add a full matrix affine constraint to the problem data. -/ def addMatrixAffineConstraint (data : ProblemData) (H : Array (Array (Array Float))) (D : Array (Array Float)) - : ProblemData := + : ProblemData := let constraint := MatrixAffine.mk D.size H D - { data with matrixAffineConstraints := + { data with matrixAffineConstraints := data.matrixAffineConstraints.push constraint } end ProblemData -end CvxLean +end CvxLean From 17f42c8961cdd82bdaa42e51059197718df4446c Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 6 Dec 2023 16:40:29 -0500 Subject: [PATCH 075/189] fix: missing case in `translateCone` --- CvxLean/Tactic/Solver/Conic.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/CvxLean/Tactic/Solver/Conic.lean b/CvxLean/Tactic/Solver/Conic.lean index cd1ae907..e7fdd0ee 100644 --- a/CvxLean/Tactic/Solver/Conic.lean +++ b/CvxLean/Tactic/Solver/Conic.lean @@ -23,6 +23,7 @@ def translateCone : ScalarConeType → CBF.ConeType | ScalarConeType.Zero => CBF.ConeType.LEq | ScalarConeType.PosOrth => CBF.ConeType.LPos | ScalarConeType.Exp => CBF.ConeType.EXP + | ScalarConeType.Q => CBF.ConeType.Q | ScalarConeType.QR => CBF.ConeType.QR -- TOOD: Change. This is hacky. From 165588753876a56a36e2d2dcbf1712ba069eb85a Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 6 Dec 2023 17:12:48 -0500 Subject: [PATCH 076/189] fix: lock output file for parallel test imports --- CvxLean/Tactic/Solver/Conic.lean | 52 +++++++++++++++----------------- 1 file changed, 25 insertions(+), 27 deletions(-) diff --git a/CvxLean/Tactic/Solver/Conic.lean b/CvxLean/Tactic/Solver/Conic.lean index e7fdd0ee..b5b7b161 100644 --- a/CvxLean/Tactic/Solver/Conic.lean +++ b/CvxLean/Tactic/Solver/Conic.lean @@ -93,40 +93,38 @@ unsafe def conicSolverFromValues (goalExprs : SolutionExpr) (data : ProblemData) (CBF.ConeProduct.mk n fixedCones.length fixedCones) -- Write input. - let inputPath := "solver/input.cbf" + let inputPath := "solver/problem.cbf" IO.FS.writeFile inputPath (ToString.toString cbf) -- Adjust path to MOSEK. let p := if let some p' := ← IO.getEnv "PATH" then - if mosekBinPath != "" then - p' ++ ":" ++ mosekBinPath - else - p' + if mosekBinPath != "" then p' ++ ":" ++ mosekBinPath else p' else mosekBinPath - -- Run solver. - let outputPath := "solver/output.sol" - let out ← IO.Process.output { - cmd := "mosek", - args := #[inputPath, "-out", outputPath], - env := #[("PATH", p)] } - if out.exitCode != 0 then - dbg_trace ("MOSEK exited with code " ++ ToString.toString out.exitCode) - return Sol.Response.failure out.exitCode.toNat - - let res := out.stdout - IO.println res - - -- Read output. - let handle ← IO.FS.Handle.mk outputPath IO.FS.Mode.read - let output ← IO.FS.Handle.readToEnd handle - - match Sol.Parser.parse output with - | Except.ok res => return Sol.Response.success res - | Except.error err => - dbg_trace ("MOSEK output parsing failed. " ++ err) - return Sol.Response.failure 1 + -- TODO: locking? + let outputPath := "solver/problem.sol" + IO.FS.withFile outputPath IO.FS.Mode.read fun handle => do + -- Run solver. + let out ← IO.Process.output { + cmd := "mosek", + args := #[inputPath], + env := #[("PATH", p)] } + if out.exitCode != 0 then + dbg_trace ("MOSEK exited with code " ++ ToString.toString out.exitCode) + return Sol.Response.failure out.exitCode.toNat + + let res := out.stdout + IO.println res + + -- Read output. + let output ← IO.FS.Handle.readToEnd handle + + match Sol.Parser.parse output with + | Except.ok res => return Sol.Response.success res + | Except.error err => + dbg_trace ("MOSEK output parsing failed. " ++ err) + return Sol.Response.failure 1 /-- TODO: Move to Generation? -/ unsafe def exprFromSol (goalExprs : SolutionExpr) (sol : Sol.Result) : MetaM Expr := do From c3d33d6587bd335161ac522d45645bb96b40317d Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 6 Dec 2023 17:14:49 -0500 Subject: [PATCH 077/189] feat: add files for case studies --- CvxLean/Examples/FittingSphere.lean | 0 CvxLean/Examples/HypersonicShapeDesign.lean | 0 CvxLean/Examples/TrajectoryOptimization.lean | 21 ++++++++++---------- CvxLean/Examples/TrussDesign.lean | 0 4 files changed, 10 insertions(+), 11 deletions(-) create mode 100644 CvxLean/Examples/FittingSphere.lean create mode 100644 CvxLean/Examples/HypersonicShapeDesign.lean create mode 100644 CvxLean/Examples/TrussDesign.lean diff --git a/CvxLean/Examples/FittingSphere.lean b/CvxLean/Examples/FittingSphere.lean new file mode 100644 index 00000000..e69de29b diff --git a/CvxLean/Examples/HypersonicShapeDesign.lean b/CvxLean/Examples/HypersonicShapeDesign.lean new file mode 100644 index 00000000..e69de29b diff --git a/CvxLean/Examples/TrajectoryOptimization.lean b/CvxLean/Examples/TrajectoryOptimization.lean index dcec1bca..79d70420 100644 --- a/CvxLean/Examples/TrajectoryOptimization.lean +++ b/CvxLean/Examples/TrajectoryOptimization.lean @@ -8,18 +8,18 @@ noncomputable section TrajectoryOptimization open Matrix -def originalBezier (n d : ℕ) - (K : Matrix (Fin (d + 2)) (Fin n) ℝ) +def originalBezier (n d : ℕ) + (K : Matrix (Fin (d + 2)) (Fin n) ℝ) (V : Matrix (Fin (d + 1)) (Fin n) ℝ) (A : Matrix (Fin d) (Fin n) ℝ) (k : Fin (d + 2) → ℝ) (v : Fin (d + 1) → ℝ) (a : Fin d → ℝ) := - optimization (x : Fin n → ℝ) (T : ℝ) - minimize T - subject to - hT : 1 ≤ T - hk : K.mulVec x ≤ k + optimization (x : Fin n → ℝ) (T : ℝ) + minimize T + subject to + hT : 1 ≤ T + hk : K.mulVec x ≤ k hv : V.mulVec x ≤ T • v ha : A.mulVec x ≤ T ^ 2 • a @@ -42,7 +42,7 @@ def convexBezier (n d : ℕ) variable {R D E : Type} [Preorder R] variable (p : Minimization D R) (q : Minimization E R) -structure Relaxation := +structure Relaxation := (r : D → E) (r_injective : Function.Injective r) (r_feasibility : ∀ x, p.constraints x → q.constraints (r x)) @@ -58,7 +58,7 @@ def relaxationBezier (n d : ℕ) Relaxation (originalBezier n d K V A k v a) (convexBezier n d K V A k v a) := { r := fun ⟨x, T⟩ => ⟨x, T, T ^ 2⟩, r_injective := fun ⟨x, T⟩ ⟨x', T'⟩ h => by rcases h with ⟨hx, hT, _⟩; rfl, - r_feasibility := fun ⟨x, T⟩ ⟨hT, hk, hv, ha⟩ => ⟨hT, hk, hv, ha, le_refl _⟩ + r_feasibility := fun ⟨x, T⟩ ⟨hT, hk, hv, ha⟩ => ⟨hT, hk, hv, ha, le_refl _⟩ r_lower_bound := fun ⟨x, T⟩ ⟨hT, _, _, _⟩ => by { simp only [convexBezier, originalBezier] rw [sqrt_le_iff] @@ -66,5 +66,4 @@ def relaxationBezier (n d : ℕ) simp [this] } } - -end TrajectoryOptimization \ No newline at end of file +end TrajectoryOptimization diff --git a/CvxLean/Examples/TrussDesign.lean b/CvxLean/Examples/TrussDesign.lean new file mode 100644 index 00000000..e69de29b From 1127666a99e99fff10fec00fff1cfd2afc887559 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 6 Dec 2023 17:44:14 -0500 Subject: [PATCH 078/189] wip: atom `logSumExp` --- .../Tactic/DCP/AtomLibrary/Fns/LogSumExp.lean | 57 +++++++++++++++++++ 1 file changed, 57 insertions(+) create mode 100644 CvxLean/Tactic/DCP/AtomLibrary/Fns/LogSumExp.lean diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogSumExp.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogSumExp.lean new file mode 100644 index 00000000..14b12e74 --- /dev/null +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogSumExp.lean @@ -0,0 +1,57 @@ +import CvxLean.Tactic.DCP.Atoms +import CvxLean.Tactic.DCP.AtomLibrary.Sets.Cones +import CvxLean.Tactic.DCP.AtomLibrary.Sets.Le +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Exp +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Log +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Sub +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Sum +import CvxLean.Lib.Math.Data.Vec + +namespace CvxLean + +open Real + +-- TODO: Move +def Vec.const (n) (x : ℝ) : Fin n → ℝ := fun _ => x + +declare_atom Vec.const [affine] (n : Nat)& (x : ℝ)? : + Vec.const n x := +bconditions +homogenity by + unfold Vec.const; ext; simp +additivity by + unfold Vec.const; ext; simp +optimality le_refl _ + +lemma Vec.sum_exp_eq_sum_div (x : Fin n → ℝ) (t : ℝ) : + Vec.sum (Vec.exp (x - Vec.const n t)) = (Vec.sum (Vec.exp x)) / (exp t) := by + unfold Vec.sum + rw [Finset.sum_div] + congr; ext i + simp [Vec.exp, Vec.const, Real.exp_sub] + +-- declare_atom logSumExp [convex] (n : ℕ)& (x : Fin n → ℝ)+ : log (Vec.sum (Vec.exp x)) := +-- bconditions +-- (h : 0 < n) +-- vconditions +-- implementationVars (t : ℝ) +-- implementationObjective t +-- implementationConstraints +-- (c1 : Vec.sum (Vec.exp (x - Vec.const n t)) ≤ 1) +-- solution (t := log (Vec.sum (Vec.exp x))) +-- solutionEqualsAtom by +-- rfl; +-- feasibility +-- (c1 : by +-- dsimp +-- simp [Vec.sum_exp_eq_sum_div, div_le_iff (Real.exp_pos _)] +-- have h : 0 < Vec.sum (Vec.exp x) := by +-- apply Finset.sum_pos +-- { intros i _; simp [Vec.exp, Real.exp_pos] } +-- { existsi 0; simp } +-- rw [Real.exp_log h]) +-- optimality by +-- intros y hy +-- rw [Vec.sum_exp_eq_sum_div] at c1 +-- sorry +-- vconditionElimination From 0f37752de924958e7d2f4f275d6ca7983c0e0080 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 10:18:23 -0500 Subject: [PATCH 079/189] feat: add bconds to feasibility --- CvxLean/Tactic/DCP/Atoms.lean | 34 ++++++++++++++++++---------------- CvxLean/Tactic/DCP/Dcp.lean | 23 +++++++++++++++-------- 2 files changed, 33 insertions(+), 24 deletions(-) diff --git a/CvxLean/Tactic/DCP/Atoms.lean b/CvxLean/Tactic/DCP/Atoms.lean index 61cddf42..772bb59d 100644 --- a/CvxLean/Tactic/DCP/Atoms.lean +++ b/CvxLean/Tactic/DCP/Atoms.lean @@ -458,23 +458,25 @@ def elabSolEqAtom (argDecls : Array LocalDecl) (vconds : Array (Lean.Name × Exp return ← mkLambdaFVars xs $ ← mkLambdaFVars cs solEqAtom /-- -/ -def elabFeas (argDecls : Array LocalDecl) (vconds : Array (Lean.Name × Expr)) +def elabFeas (argDecls : Array LocalDecl) (vconds bconds : Array (Lean.Name × Expr)) (impConstrs : Array (Lean.Name × Expr)) (sols : Array Expr) (stx : Array Syntax) : TermElabM (Array Expr) := do withExistingLocalDecls argDecls.toList do let xs := argDecls.map (mkFVar ·.fvarId) - let vconds := vconds.map fun (n,c) => (n, mkAppNBeta c xs) - let impConstrs := impConstrs.map fun (n,c) => (n, mkAppNBeta c xs) - let sols := sols.map (mkAppNBeta · xs) - withLocalDeclsDNondep vconds fun cs => do - let mut feas := #[] - for i in [:impConstrs.size] do - if (stx[i]!)[1]!.getId != impConstrs[i]!.1 then - throwError "Feasibility: Expected {impConstrs[i]!.1}, found {(stx[i]!)[1]!}" - let ty := convertLambdasToLets impConstrs[i]!.2 sols - let f ← Elab.Term.elabTermAndSynthesizeEnsuringType (stx[i]!)[3]! (some ty) - let f ← mkLambdaFVars xs $ ← mkLambdaFVars cs f - feas := feas.push f - return feas + let bconds := bconds.map fun (n,c) => (n, mkAppNBeta c xs) + withLocalDeclsDNondep bconds fun as => do + let vconds := vconds.map fun (n,c) => (n, mkAppNBeta c xs) + let impConstrs := impConstrs.map fun (n,c) => (n, mkAppNBeta c xs) + let sols := sols.map (mkAppNBeta · xs) + withLocalDeclsDNondep vconds fun cs => do + let mut feas := #[] + for i in [:impConstrs.size] do + if (stx[i]!)[1]!.getId != impConstrs[i]!.1 then + throwError "Feasibility: Expected {impConstrs[i]!.1}, found {(stx[i]!)[1]!}" + let ty := convertLambdasToLets impConstrs[i]!.2 sols + let f ← Elab.Term.elabTermAndSynthesizeEnsuringType (stx[i]!)[3]! (some ty) + let f ← mkLambdaFVars xs $ ← mkLambdaFVars as $ ← mkLambdaFVars cs f + feas := feas.push f + return feas /-- -/ def withCopyOfNonConstVars (xs : Array Expr) (argKinds : Array ArgKind) (f : Array Expr → Array Expr → TermElabM Expr) : @@ -574,7 +576,7 @@ def elabVCondElim (curv : Curvature) (argDecls : Array LocalDecl) (vconds : Arra let impConstrs ← elabImpConstrs argDecls impVars impConstrs.rawImpl let sols ← elabSols argDecls impVars impVarMap sols.rawImpl let solEqAtom ← elabSolEqAtom argDecls vconds impObj sols expr solEqAtom.raw - let feas ← elabFeas argDecls vconds impConstrs sols feas.rawImpl + let feas ← elabFeas argDecls vconds #[] impConstrs sols feas.rawImpl let opt ← elabOpt curv argDecls expr #[] impVars impObj impConstrs argKinds opt.raw let vcondElim ← elabVCondElim curv argDecls vconds vcondMap impVars impConstrs argKinds vcondElim.rawImpl @@ -638,7 +640,7 @@ def elabVCondElim (curv : Curvature) (argDecls : Array LocalDecl) (vconds : Arra let impConstrs ← elabImpConstrs argDecls impVars impConstrs.rawImpl let sols ← elabSols argDecls impVars impVarMap sols.rawImpl let solEqAtom ← elabSolEqAtom argDecls vconds impObj sols expr solEqAtom.raw - let feas ← elabFeas argDecls vconds impConstrs sols feas.rawImpl + let feas ← elabFeas argDecls vconds bconds impConstrs sols feas.rawImpl let opt ← elabOpt curv argDecls expr bconds impVars impObj impConstrs argKinds opt.raw let vcondElim ← elabVCondElim curv argDecls vconds vcondMap impVars impConstrs argKinds vcondElim.rawImpl diff --git a/CvxLean/Tactic/DCP/Dcp.lean b/CvxLean/Tactic/DCP/Dcp.lean index ade0c6ae..f67785fc 100644 --- a/CvxLean/Tactic/DCP/Dcp.lean +++ b/CvxLean/Tactic/DCP/Dcp.lean @@ -367,17 +367,19 @@ abbrev FeasibilityProofsTree := Tree FeasibilityProofs Unit partial def mkFeasibility : GraphAtomDataTree → ReducedExprsWithSolutionTree → + BCondsTree → VCondsTree → SolEqAtomProofsTree → MetaM FeasibilityProofsTree | Tree.node atom childAtoms, Tree.node _reducedWithSolution childReducedWithSolution, + Tree.node bconds childBConds, Tree.node vcondVars childVCondVars, Tree.node _solEqAtom childSolEqAtom => do -- Recursive calls for arguments. let mut childFeasibility := #[] for i in [:childAtoms.size] do - let c ← mkFeasibility childAtoms[i]! childReducedWithSolution[i]! childVCondVars[i]! childSolEqAtom[i]! + let c ← mkFeasibility childAtoms[i]! childReducedWithSolution[i]! childBConds[i]! childVCondVars[i]! childSolEqAtom[i]! childFeasibility := childFeasibility.push c -- Use solEqAtom of children to rewrite the arguments in the vconditions. let mut vconds := #[] @@ -389,12 +391,13 @@ partial def mkFeasibility : -- Apply feasibility property of the atom. let feasibility := atom.feasibility let feasibility := feasibility.map (mkAppN · (childReducedWithSolution.map (·.val.1))) + let feasibility := feasibility.map (mkAppN · bconds) let feasibility := feasibility.map (mkAppN · vconds) let _ ← feasibility.mapM check return Tree.node feasibility childFeasibility - | Tree.leaf _, Tree.leaf _, Tree.leaf _, Tree.leaf _ => + | Tree.leaf _, Tree.leaf _, Tree.leaf _, Tree.leaf _, Tree.leaf _ => return Tree.leaf () - | _, _, _, _ => + | _, _, _, _, _ => throwError "Tree mismatch" abbrev OptimalityProof := Expr @@ -747,12 +750,16 @@ withExistingLocalDecls originalVarsDecls.toList do /-- -/ def mkFeasibilityOC (originalVarsDecls : Array LocalDecl) - (atoms : OC (Tree GraphAtomData Expr)) (reducedWithSolution : OC ReducedExprsWithSolutionTree) - (vcondVars : OC (Tree (Array Expr) Unit)) (originalConstrVars : Array LocalDecl) (solEqAtom : OC (Tree Expr Expr)) : - MetaM (OC (Tree (Array Expr) Unit)) := + (atoms : OC GraphAtomDataTree) + (reducedWithSolution : OC ReducedExprsWithSolutionTree) + (bconds : OC BCondsTree) + (vcondVars : OC VCondsTree) + (originalConstrVars : Array LocalDecl) + (solEqAtom : OC SolEqAtomProofsTree) : + MetaM (OC FeasibilityProofsTree) := withExistingLocalDecls originalVarsDecls.toList do withExistingLocalDecls originalConstrVars.toList do - let feasibility ← OC.map4M mkFeasibility atoms reducedWithSolution vcondVars solEqAtom + let feasibility ← OC.map5M mkFeasibility atoms reducedWithSolution bconds vcondVars solEqAtom trace[Meta.debug] "feasibility {feasibility}" return feasibility /-- -/ @@ -844,7 +851,7 @@ def mkProcessedAtomTree (objCurv : Curvature) (objFun : Expr) (constraints : Lis let forwardImagesNewVars ← withExistingLocalDecls originalVarsDecls.toList do mkForwardImagesNewVars reducedWithSolution let solEqAtom ← mkSolEqAtomOC originalVarsDecls atoms reducedWithSolution vcondVars originalConstrVars - let feasibility ← mkFeasibilityOC originalVarsDecls atoms reducedWithSolution vcondVars + let feasibility ← mkFeasibilityOC originalVarsDecls atoms reducedWithSolution bconds vcondVars originalConstrVars solEqAtom let reducedExprs ← mkReducedExprsOC originalVarsDecls newVarDecls atoms newVars let (newConstrs, newConstrVars, newConstrVarsArray) From 8531920c9d3122098f911eb76f2163f4e669d87a Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 10:59:35 -0500 Subject: [PATCH 080/189] fix: build feasibility expression with bconds --- CvxLean/Tactic/DCP/Atoms.lean | 104 ++++++++++++++++------------------ 1 file changed, 49 insertions(+), 55 deletions(-) diff --git a/CvxLean/Tactic/DCP/Atoms.lean b/CvxLean/Tactic/DCP/Atoms.lean index 772bb59d..bd9d0550 100644 --- a/CvxLean/Tactic/DCP/Atoms.lean +++ b/CvxLean/Tactic/DCP/Atoms.lean @@ -189,6 +189,7 @@ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandEla -- trace[Meta.debug] "a: {←inferType a}")) let vconds := atomData.vconds.map fun (n,c) => (n, mkAppNBeta c xs) + let bconds := atomData.bconds.map fun (n,c) => (n, mkAppNBeta c xs) let solEqAtomProofs := pat.solEqAtom.constr.map Tree.val @@ -204,10 +205,12 @@ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandEla let feasXs := mkAppNBeta feas xs let adjustedFeas ← withLocalDeclsDNondep vconds fun cs => do - let feasXsVconds := mkAppNBeta feasXs cs - let proofAdjusted := solEqAtomProofs[i]!.replaceFVars vs1 vs1Sol - let adjustedFeas ← mkAppM ``Eq.mpr #[proofAdjusted, feasXsVconds] - mkLambdaFVars xs <| ← mkLambdaFVars cs adjustedFeas + withLocalDeclsDNondep bconds fun bs => do + let feasXsVconds := mkAppNBeta feasXs cs + let feasXsConds := mkAppNBeta feasXsVconds bs + let proofAdjusted := solEqAtomProofs[i]!.replaceFVars vs1 vs1Sol + let adjustedFeas ← mkAppM ``Eq.mpr #[proofAdjusted, feasXsConds] + mkLambdaFVars (xs ++ cs ++ bs) adjustedFeas oldFeasibilityAdjusted := oldFeasibilityAdjusted.push adjustedFeas for proof in solEqAtomProofs do @@ -218,7 +221,8 @@ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandEla let newFeasibility ← newConstrProofs.mapM (fun e => withLocalDeclsDNondep vconds fun cs => do - mkLambdaFVars xs <| ← mkLambdaFVars cs (e.replaceFVars vs1 vs1Sol)) + withLocalDeclsDNondep bconds fun bs => do + mkLambdaFVars (xs ++ cs ++ bs) (e.replaceFVars vs1 vs1Sol)) for f in atomData.feasibility do trace[Meta.debug] "feasibility: {← inferType f}" @@ -240,57 +244,47 @@ def reduceAtomData (objCurv : Curvature) (atomData : GraphAtomData) : CommandEla trace[Meta.debug] "objFunFromReducedObjFun: {← inferType objFunFromReducedObjFun}" trace[Meta.debug] "pat.optimality.objFun: {← inferType atomData.optimality}" - let optimalityXs := mkAppN atomData.optimality (xs ++ vs1) - trace[Meta.debug] "newOptimality: {← inferType optimalityXs}" + let newOptimality ← - withLocalDeclsDNondep (reducedConstrs.map (fun rc => (`_, rc))) fun cs => do - -- First, apply all constraints. - let mut optimalityAfterReduced := optimalityXs - for i in [:reducedConstrs.size] do - let c := mkApp constraintsFromReducedConstraints[i]! cs[i]! - optimalityAfterReduced := mkApp optimalityAfterReduced c - -- Then, adjust the condition using `objFunFromReducedObjFun`. - trace[Meta.debug] "optimalityAfterReduced: {← inferType optimalityAfterReduced}" - let monoArgsCount := getMonoArgsCount objCurv atomData.argKinds - let optimalityAfterApplyWithConditionAdjusted ← - lambdaTelescope (← whnf optimalityAfterReduced) <| fun xs e => do - -- Every extra argument has an extra condition, e.g. x', x ≤ x. - let monoArgs := xs[:2 * monoArgsCount] - let optCondition ← mkLambdaFVars xs[2 * monoArgsCount:] e - let newCond ← - if atomData.curvature == Curvature.Convex then - mkAppOptM ``le_trans #[ - atomRange, none, none, none, none, - optCondition, objFunFromReducedObjFun] - else - -- TODO: concave. but convex_set too? - mkAppOptM ``le_trans #[ - atomRange, none, none, none, none, - objFunFromReducedObjFun, optCondition] - mkLambdaFVars monoArgs newCond - - trace[Meta.debug] "optimalityAfterApplyWithConditionAdjusted: {← inferType optimalityAfterApplyWithConditionAdjusted}" - trace[Meta.debug] "newOptimality applied: {← inferType optimalityAfterReduced}" - let ds := pat.newConstrVarsArray.map (mkFVar ·.fvarId) - mkLambdaFVars (xs ++ vs1 ++ vs2) <| ← mkLambdaFVars (cs ++ ds) optimalityAfterApplyWithConditionAdjusted - -- trace[Meta.debug] "newOptimality2: {← inferType newOptimality2}" - - -- -- NOTE: new optimality might be of the form ∀ i. x i ≤ y i, so - -- -- we cannot apply le_trans directly. - -- let x ← lambdaTelescope (← whnf newOptimality2) (fun xs a => do - -- trace[Meta.debug] "a: {← inferType a}" - -- trace[Meta.debug] "xs: {xs}" - -- return ()) - -- match (← inferType objFunFromReducedObjFun).le? with - -- | some (t, y, z) => - -- trace[Meta.debug] "atomData.argKinds: {atomData.argKinds}" - -- let x ← getOptArgs objCurv xs atomData.argKinds - -- trace[Meta.debug] "x: {x}" - -- -- mkAppOptM ``le_trans #[ - -- -- t, none, none, y, z, - -- -- newOptimality2, objFunFromReducedObjFun] - -- let ds := pat.newConstrVarsArray.map (mkFVar ·.fvarId) - -- mkLambdaFVars (xs ++ vs1 ++ vs2) <| ← mkLambdaFVars (cs ++ ds) newOptimality2 --optimalityAdjusted + withLocalDeclsDNondep bconds fun bs => do + let optimalityXsBConds := mkAppN atomData.optimality (xs ++ bs ++ vs1) + trace[Meta.debug] "newOptimality: {← inferType optimalityXsBConds}" + withLocalDeclsDNondep (reducedConstrs.map (fun rc => (`_, rc))) fun cs => do + -- First, apply all constraints. + let mut optimalityAfterReduced := optimalityXsBConds + for i in [:reducedConstrs.size] do + let c := mkApp constraintsFromReducedConstraints[i]! cs[i]! + optimalityAfterReduced := mkApp optimalityAfterReduced c + -- Then, adjust the condition using `objFunFromReducedObjFun`. + trace[Meta.debug] "optimalityAfterReduced: {← inferType optimalityAfterReduced}" + let monoArgsCount := getMonoArgsCount objCurv atomData.argKinds + let optimalityAfterApplyWithConditionAdjusted ← + lambdaTelescope (← whnf optimalityAfterReduced) <| fun xs e => do + -- Every extra argument has an extra condition, e.g. x', x ≤ x. + trace[Meta.debug] "xs: {xs}" + let monoArgs := xs[:2 * monoArgsCount] + trace[Meta.debug] "monoArgs: {monoArgs}" + trace[Meta.debug] "e: {← inferType e}" + let optCondition ← mkLambdaFVars xs[2 * monoArgsCount:] e + let newCond ← + if atomData.curvature == Curvature.Convex then + mkAppOptM ``le_trans #[ + atomRange, none, none, none, none, + optCondition, objFunFromReducedObjFun] + else + -- TODO: concave. but convex_set too? + mkAppOptM ``le_trans #[ + atomRange, none, none, none, none, + objFunFromReducedObjFun, optCondition] + mkLambdaFVars monoArgs newCond + + trace[Meta.debug] "optimalityAfterApplyWithConditionAdjusted: {← inferType optimalityAfterApplyWithConditionAdjusted}" + trace[Meta.debug] "newOptimality applied: {← inferType optimalityAfterReduced}" + let ds := pat.newConstrVarsArray.map (mkFVar ·.fvarId) + mkLambdaFVars (xs ++ bs ++ vs1 ++ vs2 ++ cs ++ ds) optimalityAfterApplyWithConditionAdjusted + + trace[Meta.debug] "newOptimality: {← inferType newOptimality}" + let mut newVCondElims := #[] for vcondElim in atomData.vcondElim do let newVCondElim := mkAppN vcondElim (xs ++ vs1) From 410a6546bd1f2b66f307b13578f502a37ab51ae8 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 10:59:45 -0500 Subject: [PATCH 081/189] feat: general `logSumExp` --- .../Tactic/DCP/AtomLibrary/Fns/LogSumExp.lean | 56 ++++++++++--------- 1 file changed, 31 insertions(+), 25 deletions(-) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogSumExp.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogSumExp.lean index 14b12e74..58194221 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogSumExp.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogSumExp.lean @@ -30,28 +30,34 @@ lemma Vec.sum_exp_eq_sum_div (x : Fin n → ℝ) (t : ℝ) : congr; ext i simp [Vec.exp, Vec.const, Real.exp_sub] --- declare_atom logSumExp [convex] (n : ℕ)& (x : Fin n → ℝ)+ : log (Vec.sum (Vec.exp x)) := --- bconditions --- (h : 0 < n) --- vconditions --- implementationVars (t : ℝ) --- implementationObjective t --- implementationConstraints --- (c1 : Vec.sum (Vec.exp (x - Vec.const n t)) ≤ 1) --- solution (t := log (Vec.sum (Vec.exp x))) --- solutionEqualsAtom by --- rfl; --- feasibility --- (c1 : by --- dsimp --- simp [Vec.sum_exp_eq_sum_div, div_le_iff (Real.exp_pos _)] --- have h : 0 < Vec.sum (Vec.exp x) := by --- apply Finset.sum_pos --- { intros i _; simp [Vec.exp, Real.exp_pos] } --- { existsi 0; simp } --- rw [Real.exp_log h]) --- optimality by --- intros y hy --- rw [Vec.sum_exp_eq_sum_div] at c1 --- sorry --- vconditionElimination +lemma Vec.sum_exp_pos {n} (hn : 0 < n) (x : Fin n → ℝ) : + 0 < Vec.sum (Vec.exp x) := by + apply Finset.sum_pos + { intros i _; simp [Vec.exp, Real.exp_pos] } + { existsi ⟨0, hn⟩; simp } + +declare_atom logSumExp [convex] (n : ℕ)& (x : Fin n → ℝ)+ : log (Vec.sum (Vec.exp x)) := +bconditions + (hn : 0 < n) +vconditions +implementationVars (t : ℝ) +implementationObjective t +implementationConstraints + (c1 : Vec.sum (Vec.exp (x - Vec.const n t)) ≤ 1) +solution (t := log (Vec.sum (Vec.exp x))) +solutionEqualsAtom by + rfl; +feasibility + (c1 : by + dsimp + simp [Vec.sum_exp_eq_sum_div, div_le_iff (Real.exp_pos _)] + simp [Real.exp_log (Vec.sum_exp_pos hn x)]) +optimality by + intros y hy + simp [Vec.sum_exp_eq_sum_div, div_le_iff (exp_pos _)] at c1 + rw [←log_exp t, log_le_log (Vec.sum_exp_pos hn y) (exp_pos _)] + refine le_trans ?_ c1 + apply Finset.sum_le_sum + intros i _ + simp [Vec.exp, hy i] +vconditionElimination From f97a75a8eb8d1c666052fdb275ecab98e9634151 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 11:08:44 -0500 Subject: [PATCH 082/189] chore: remove old matrix lemmas --- CvxLean/Lib/Math/Data/Matrix.lean | 194 +++--------------------------- 1 file changed, 17 insertions(+), 177 deletions(-) diff --git a/CvxLean/Lib/Math/Data/Matrix.lean b/CvxLean/Lib/Math/Data/Matrix.lean index 48cd948f..89ac1855 100644 --- a/CvxLean/Lib/Math/Data/Matrix.lean +++ b/CvxLean/Lib/Math/Data/Matrix.lean @@ -8,23 +8,17 @@ import CvxLean.Lib.Math.Data.List namespace Matrix --- Order - instance [Preorder α] : Preorder (Matrix m n α) := { le := fun A B => ∀ i j, A i j ≤ B i j le_refl := fun _ _ _ => le_refl _ le_trans := fun _ _ _ hAB hBC i j => le_trans (hAB i j) (hBC i j) lt_iff_le_not_le := fun _ _ => refl _ } --- Abs - def abs (A : Matrix m n ℝ) : Matrix m n ℝ := fun i j => Abs.abs (A i j) instance : Abs (Matrix m n ℝ) := ⟨abs⟩ --- Vec cons - theorem vecCons_zero_zero {n} : vecCons (0 : ℝ) (0 : Fin n → ℝ) = 0 := by ext i ; refine' Fin.cases _ _ i <;> simp [vecCons] @@ -37,149 +31,12 @@ theorem add_vecCons {n} (x : ℝ) (v : Fin n → ℝ) (y : ℝ) (w : Fin n → : vecCons x v + vecCons y w = vecCons (x + y) (v + w) := by ext i ; refine' Fin.cases _ _ i <;> simp [vecCons] --- Sum - open BigOperators def sum [Fintype m] [AddCommMonoid α] (X : Matrix m m α) : α := ∑ i, (∑ j, X i j) --- theorem vecCons_zero_zero {n} --- : Matrix.vecCons (0 : ℝ) (0 : Fin n → ℝ) = 0 := by --- ext i ; refine' Fin.cases _ _ i <;> simp [Matrix.vecCons] - --- theorem smul_vecCons {n} (x : ℝ) (y : ℝ) (v : Fin n → ℝ) --- : x • Matrix.vecCons y v = Matrix.vecCons (x • y) (x • v) := by --- ext i ; refine' Fin.cases _ _ i <;> simp [Matrix.vecCons] - --- theorem add_vecCons {n} (x : ℝ) (v : Fin n → ℝ) (y : ℝ) (w : Fin n → ℝ) --- : Matrix.vecCons x v + Matrix.vecCons y w = Matrix.vecCons (x + y) (v + w) := by --- ext i ; refine' Fin.cases _ _ i <;> simp [Matrix.vecCons] - --- theorem dotProduct_zero'' {m} [Fintype m] (x : m → ℝ) --- : Matrix.dotProduct x 0 = 0 := by --- simp [Matrix.dotProduct] - --- theorem dotProduct_smul' {m} [Fintype m] (x : m → ℝ) (y : m → ℝ) (a : ℝ) --- : Matrix.dotProduct x (a • y) = a • Matrix.dotProduct x y := by --- unfold Matrix.dotProduct ; rw [Finset.smul_sum] --- apply congr_arg ; ext i ; simp ; ring - --- theorem smul_dotProduct' {m} [Fintype m] (x : m → ℝ) (y : m → ℝ) (a : ℝ) --- : Matrix.dotProduct (a • x) y = a • Matrix.dotProduct x y := by --- unfold Matrix.dotProduct ; rw [Finset.smul_sum] --- apply congr_arg ; ext i ; simp ; ring - --- theorem dotProduct_add' {m} [Fintype m] (x : m → ℝ) (y y' : m → ℝ) --- : Matrix.dotProduct x (y + y') = Matrix.dotProduct x y + Matrix.dotProduct x y' := by --- unfold Matrix.dotProduct ; simp only [←Finset.sum_add_distrib] --- apply congr_arg ; ext i ; simp ; ring - --- theorem add_dotProduct' {m} [Fintype m] (x x' : m → ℝ) (y : m → ℝ) --- : Matrix.dotProduct (x + x') y = Matrix.dotProduct x y + Matrix.dotProduct x' y := by --- unfold Matrix.dotProduct ; simp only [←Finset.sum_add_distrib] --- apply congr_arg ; ext i ; simp ; ring - --- theorem sum_zero {n} --- : Matrix.sum (0 : Matrix (Fin n) (Fin n) ℝ) = 0 := by --- simp [Matrix.sum, Matrix.zero_apply] - --- theorem smul_sum {n} (X : Matrix (Fin n) (Fin n) ℝ) (κ : ℝ) --- : κ • Matrix.sum X = Matrix.sum (κ • X) := by --- unfold Matrix.sum ; rw [Finset.smul_sum] --- congr ; ext i ; rw [Finset.smul_sum] ; rfl - --- theorem sum_add {n} (X Y : Matrix (Fin n) (Fin n) ℝ) --- : Matrix.sum X + Matrix.sum Y = Matrix.sum (X + Y) := by --- unfold Matrix.sum ; rw [←Finset.sum_add_distrib] --- congr ; ext i ; rw [←Finset.sum_add_distrib] ; rfl - --- theorem diag_smul' {m} [Fintype m] (x : ℝ) (A : Matrix m m ℝ) --- : Matrix.diag (x • A) = x • Matrix.diag A := by --- rfl - --- theorem diagonal_zero' {n} --- : Matrix.diagonal (0 : Fin n → ℝ) = 0 := by --- funext i j ; by_cases i = j <;> --- simp [Matrix.diagonal, h] - --- theorem diagonal_smul' {n} (d : Fin n → ℝ) (κ : ℝ) --- : κ • Matrix.diagonal d = Matrix.diagonal (κ • d) := by --- funext i j ; by_cases i = j <;> --- simp [Matrix.diagonal, h] - --- theorem diagonal_add' {n} (d₁ d₂ : Fin n → ℝ) --- : Matrix.diagonal d₁ + Matrix.diagonal d₂ = Matrix.diagonal (d₁ + d₂) := by --- funext i j ; by_cases i = j <;> --- simp [Matrix.diagonal, h] - --- theorem trace_zero' {m} [Fintype m] --- : Matrix.trace (0 : Matrix m m ℝ) = 0 := by --- unfold Matrix.trace ; rw [Matrix.diag_zero] --- exact Finset.sum_const_zero - --- theorem trace_smul' {m} [Fintype m] (A : Matrix m m ℝ) (κ : ℝ) --- : Matrix.trace (κ • A) = κ • Matrix.trace A := by --- unfold Matrix.trace ; rw [Matrix.diag_smul', Finset.smul_sum] ; rfl - --- theorem trace_add' {m} [Fintype m] (A B : Matrix m m ℝ) --- : Matrix.trace (A + B) = Matrix.trace A + Matrix.trace B := by --- unfold Matrix.trace ; rw [Matrix.diag_add, ←Finset.sum_add_distrib] ; rfl - --- theorem transpose_zero' {m} [Fintype m] --- : Matrix.transpose (0 : Matrix m m ℝ) = 0 := by --- funext i _ ; simp [Matrix.transpose, id] - --- theorem transpose_add' {m} [Fintype m] (A B : Matrix m m ℝ) --- : Matrix.transpose (A + B) = Matrix.transpose A + Matrix.transpose B := by --- funext i j ; simp [Matrix.transpose, id] - --- theorem fromBlocks_smul' {n m l o} (κ : ℝ) --- (A : Matrix n l ℝ) (B : Matrix n m ℝ) (C : Matrix o l ℝ) (D : Matrix o m ℝ) --- : κ • Matrix.fromBlocks A B C D = Matrix.fromBlocks (κ • A) (κ • B) (κ • C) (κ • D) := by --- funext i j ; cases i <;> cases j <;> rw [Pi.smul_apply, Pi.smul_apply] <;> --- simp [Matrix.fromBlocks] - --- theorem mul_zero' {m} [Fintype m] (A : Matrix m m ℝ) --- : A * (0 : Matrix m m ℝ) = 0 := by --- funext ; exact Matrix.dotProduct_zero'' _ - --- theorem zero_mul' {m} [Fintype m] (A : Matrix m m ℝ) --- : (0 : Matrix m m ℝ) * A = 0 := by --- funext ; exact Matrix.zero_dotProduct' _ - --- theorem mul_smul' {m} [Fintype m] (κ : ℝ) (A : Matrix m m ℝ) (B : Matrix m m ℝ) --- : A * (κ • B) = κ • (A * B) := by --- funext ; exact Matrix.dotProduct_smul' _ _ _ - --- theorem smul_mul' {m} [Fintype m] (κ : ℝ) (A : Matrix m m ℝ) (B : Matrix m m ℝ) --- : (κ • A) * B = κ • (A * B) := by --- funext ; exact Matrix.smul_dotProduct' _ _ _ - --- theorem mul_add' {m} [Fintype m] (A : Matrix m m ℝ) (B C : Matrix m m ℝ) --- : A * (B + C) = (A * B) + (A * C) := by --- funext ; exact Matrix.dotProduct_add' _ _ _ - --- theorem add_mul' {m} [Fintype m] (A B : Matrix m m ℝ) (C : Matrix m m ℝ) --- : (A + B) * C = (A * C) + (B * C) := by --- funext ; exact Matrix.add_dotProduct' _ _ _ - --- theorem mulVec_zero' {m n} [Fintype m] [Fintype n] (A : Matrix m n ℝ) --- : Matrix.mulVec A (0 : n → ℝ) = 0 := by --- funext i ; unfold Matrix.mulVec --- convert @Matrix.dotProduct_zero'' n _ (fun j => A i j) - --- theorem mulVec_smul' {m n} [Fintype m] [Fintype n] --- (κ : ℝ) (A : Matrix m n ℝ) (v : n → ℝ) --- : Matrix.mulVec A (κ • v) = κ • Matrix.mulVec A v := by --- funext i ; simp [Matrix.mulVec] - --- theorem mulVec_add' {m n} [Fintype m] [Fintype n] --- (A : Matrix m n ℝ) (v w : n → ℝ) --- : Matrix.mulVec A (v + w) = Matrix.mulVec A v + Matrix.mulVec A w := by --- funext i ; simp [Matrix.mulVec] - --- Transform to arrays to compute. Avoiding mathbin matrix operations. +-- Transform to arrays to compute. namespace Computable variable {α} [Zero α] @@ -197,10 +54,7 @@ def toArray (M : Matrix (Fin n) (Fin n) α) : Array (Array α) := def dotProduct [Mul α] [Add α] (v w : Fin n → α) : α := ((Array.zip (vecToArray v) (vecToArray w)).map (fun ⟨a, b⟩ => a * b)).foldl (· + ·) 0 --- TODO: temporary hack because mathbin breaks infixl --- infixl:72 " ⬝ᵥᶜ " => Matrix.Computable.dotProduct -macro:72 l:term:72 " ⬝ᵥᶜ " r:term:73 : term => - `(Matrix.Computable.dotProduct $l $r) +infixl:72 " ⬝ᵥᶜ " => Matrix.Computable.dotProduct def mulVec [Mul α] [Add α] (M : Matrix (Fin m) (Fin n) α) (v : (Fin n) → α) : Fin m → α | i => (fun j => M i j) ⬝ᵥᶜ v @@ -214,28 +68,14 @@ def transpose (α) (M : Matrix m n α) : Matrix n m α def diag (α) (M : Matrix n n α) : n → α | x => M x x --- instance [Add α] : Add (Matrix m n α) where --- add A B i j := (A i j) + (B i j) - --- instance [Sub α] : Sub (Matrix m n α) where --- sub A B i j := (A i j) - (B i j) - --- instance [Abs α] : Abs (Matrix m n α) where --- abs A i j := abs (A i j) - def mul [Mul α] [Add α] (M : Matrix (Fin l) (Fin m) α) (N : Matrix (Fin m) (Fin n) α) : Matrix (Fin l) (Fin n) α := fun i k => (fun j => M i j) ⬝ᵥᶜ (fun j => N j k) --- TODO: temporary hack because mathbin breaks infixl --- infixl:75 " ⬝ᶜ " => Matrix.Computable.mul -macro:75 l:term " ⬝ᶜ " r:term : term => - `(Matrix.Computable.mul $l $r) +infixl:75 " ⬝ᶜ " => Matrix.Computable.mul def tr [Add α] (A : Matrix (Fin n) (Fin n) α) : α := (Computable.vecToArray (fun i => A i i)).foldl (· + ·) 0 --- Project-specific. - def PosDef [Add α] [Mul α] [LT α] (A : Matrix (Fin n) (Fin n) α) : Prop := ∀ x : (Fin n) → α, Matrix.Computable.vecMul x A ⬝ᵥᶜ x > 0 @@ -243,30 +83,30 @@ def PosSemidef [Add α] [Mul α] [LE α] (A : Matrix (Fin n) (Fin n) α) : Prop ∀ x : (Fin n) → α, Matrix.Computable.vecMul x A ⬝ᵥᶜ x ≥ 0 def posDefObjective [Add α] [Mul α] (C X : Matrix (Fin n) (Fin n) α) : α := -Matrix.Computable.tr (Matrix.Computable.mul C X) + Matrix.Computable.tr (Matrix.Computable.mul C X) def quadForm [Add α] [Mul α] (A : Matrix (Fin n) (Fin n) α) (x : (Fin n) → α) := -Matrix.Computable.vecMul x A ⬝ᵥᶜ x + Matrix.Computable.vecMul x A ⬝ᵥᶜ x def covarianceMatrix [Add α] [Mul α] [Div α] {N n : ℕ} [OfNat α N] (Y : Matrix (Fin N) (Fin n) α) (i j : Fin n) := -((Computable.vecToArray Y).map (fun y => (y i) * y j)).foldl (· + ·) 0 / (OfNat.ofNat N) + ((Computable.vecToArray Y).map (fun y => (y i) * y j)).foldl (· + ·) 0 / (OfNat.ofNat N) def diagonal {n : Type u_3} {α : Type v} [DecidableEq n] [Zero α] (x : n → α) : Matrix n n α := -fun i j => (if i = j then x i else 0) + fun i j => (if i = j then x i else 0) def fromBlocks {l : Type} {m : Type} {n : Type} {o : Type} {α : Type} : Matrix n l α → Matrix n m α → Matrix o l α → Matrix o m α → Matrix (Sum n o) (Sum l m) α := -fun A B C D i j => by - cases i with - | inl i => - cases j with - | inl j => exact A i j - | inr j => exact B i j - | inr i => - cases j with - | inl j => exact C i j - | inr j => exact D i j + fun A B C D i j => by + cases i with + | inl i => + cases j with + | inl j => exact A i j + | inr j => exact B i j + | inr i => + cases j with + | inl j => exact C i j + | inr j => exact D i j end Computable From fdd4d697fe31ab603c616d026cf3f3843edaea24 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 11:09:07 -0500 Subject: [PATCH 083/189] refactor: move `Vec.const` --- CvxLean/Lib/Math/Data/Vec.lean | 15 +++++++++------ .../Tactic/DCP/AtomLibrary/Fns/VecConst.lean | 17 +++++++++++++++++ 2 files changed, 26 insertions(+), 6 deletions(-) create mode 100644 CvxLean/Tactic/DCP/AtomLibrary/Fns/VecConst.lean diff --git a/CvxLean/Lib/Math/Data/Vec.lean b/CvxLean/Lib/Math/Data/Vec.lean index 95fa724e..c530fb68 100644 --- a/CvxLean/Lib/Math/Data/Vec.lean +++ b/CvxLean/Lib/Math/Data/Vec.lean @@ -27,21 +27,24 @@ open Real BigOperators instance : Norm (m → ℝ) where norm x := sqrt (∑ i, (x i) ^ 2) -variable (x y : m → Real) +variable (x y : m → ℝ) -def exp : m → Real := +def const (n : ℕ) (k : ℝ) : Fin n → ℝ := + fun _ => k + +def exp : m → ℝ := fun i => Real.exp (x i) -def log : m → Real := +def log : m → ℝ := fun i => Real.log (x i) -def entr : m → Real := +def entr : m → ℝ := fun i => Real.entr (x i) -def huber : m → Real := +def huber : m → ℝ := fun i => Real.huber (x i) -def klDiv : m → Real := +def klDiv : m → ℝ := fun i => Real.klDiv (x i) (y i) end Real diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/VecConst.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/VecConst.lean new file mode 100644 index 00000000..bd6708d4 --- /dev/null +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/VecConst.lean @@ -0,0 +1,17 @@ +import CvxLean.Tactic.DCP.Atoms +import CvxLean.Lib.Math.Data.Vec + +namespace CvxLean + +open Real + +declare_atom Vec.const [affine] (n : Nat)& (x : ℝ)? : + Vec.const n x := +bconditions +homogenity by + unfold Vec.const; ext; simp +additivity by + unfold Vec.const; ext; simp +optimality le_refl _ + +end CvxLean From b92dd1e6c35c8e9f07ae87007053e9703e8a70b8 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 11:15:55 -0500 Subject: [PATCH 084/189] feat: `logSumExp` split and version with 2 elements --- CvxLean/Lib/Math/LogSumExp.lean | 21 +++++++ .../Tactic/DCP/AtomLibrary/Fns/LogSumExp.lean | 59 +++++++++---------- 2 files changed, 50 insertions(+), 30 deletions(-) create mode 100644 CvxLean/Lib/Math/LogSumExp.lean diff --git a/CvxLean/Lib/Math/LogSumExp.lean b/CvxLean/Lib/Math/LogSumExp.lean new file mode 100644 index 00000000..011b9b37 --- /dev/null +++ b/CvxLean/Lib/Math/LogSumExp.lean @@ -0,0 +1,21 @@ +import CvxLean.Lib.Math.Data.Real +import CvxLean.Lib.Math.Data.Vec + +namespace Vec + +open Real + +lemma sum_exp_eq_sum_div (x : Fin n → ℝ) (t : ℝ) : + Vec.sum (Vec.exp (x - Vec.const n t)) = (Vec.sum (Vec.exp x)) / (Real.exp t) := by + unfold Vec.sum + rw [Finset.sum_div] + congr; ext i + simp [Vec.exp, Vec.const, Real.exp_sub] + +lemma sum_exp_pos {n} (hn : 0 < n) (x : Fin n → ℝ) : + 0 < Vec.sum (Vec.exp x) := by + apply Finset.sum_pos + { intros i _; simp [Vec.exp, Real.exp_pos] } + { existsi ⟨0, hn⟩; simp } + +end Vec diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogSumExp.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogSumExp.lean index 58194221..b98dbc72 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogSumExp.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogSumExp.lean @@ -5,37 +5,14 @@ import CvxLean.Tactic.DCP.AtomLibrary.Fns.Exp import CvxLean.Tactic.DCP.AtomLibrary.Fns.Log import CvxLean.Tactic.DCP.AtomLibrary.Fns.Sub import CvxLean.Tactic.DCP.AtomLibrary.Fns.Sum -import CvxLean.Lib.Math.Data.Vec +import CvxLean.Tactic.DCP.AtomLibrary.Fns.VecCons +import CvxLean.Tactic.DCP.AtomLibrary.Fns.VecConst +import CvxLean.Lib.Math.LogSumExp namespace CvxLean open Real --- TODO: Move -def Vec.const (n) (x : ℝ) : Fin n → ℝ := fun _ => x - -declare_atom Vec.const [affine] (n : Nat)& (x : ℝ)? : - Vec.const n x := -bconditions -homogenity by - unfold Vec.const; ext; simp -additivity by - unfold Vec.const; ext; simp -optimality le_refl _ - -lemma Vec.sum_exp_eq_sum_div (x : Fin n → ℝ) (t : ℝ) : - Vec.sum (Vec.exp (x - Vec.const n t)) = (Vec.sum (Vec.exp x)) / (exp t) := by - unfold Vec.sum - rw [Finset.sum_div] - congr; ext i - simp [Vec.exp, Vec.const, Real.exp_sub] - -lemma Vec.sum_exp_pos {n} (hn : 0 < n) (x : Fin n → ℝ) : - 0 < Vec.sum (Vec.exp x) := by - apply Finset.sum_pos - { intros i _; simp [Vec.exp, Real.exp_pos] } - { existsi ⟨0, hn⟩; simp } - declare_atom logSumExp [convex] (n : ℕ)& (x : Fin n → ℝ)+ : log (Vec.sum (Vec.exp x)) := bconditions (hn : 0 < n) @@ -43,21 +20,43 @@ vconditions implementationVars (t : ℝ) implementationObjective t implementationConstraints - (c1 : Vec.sum (Vec.exp (x - Vec.const n t)) ≤ 1) + (c : Vec.sum (Vec.exp (x - Vec.const n t)) ≤ 1) solution (t := log (Vec.sum (Vec.exp x))) solutionEqualsAtom by rfl; feasibility - (c1 : by + (c : by dsimp simp [Vec.sum_exp_eq_sum_div, div_le_iff (Real.exp_pos _)] simp [Real.exp_log (Vec.sum_exp_pos hn x)]) optimality by intros y hy - simp [Vec.sum_exp_eq_sum_div, div_le_iff (exp_pos _)] at c1 + simp [Vec.sum_exp_eq_sum_div, div_le_iff (exp_pos _)] at c rw [←log_exp t, log_le_log (Vec.sum_exp_pos hn y) (exp_pos _)] - refine le_trans ?_ c1 + refine le_trans ?_ c apply Finset.sum_le_sum intros i _ simp [Vec.exp, hy i] vconditionElimination + +declare_atom logSumExp₂ [convex] (x : ℝ)+ (y : ℝ)+ : log ((exp x) + (exp y)) := +vconditions +implementationVars (t : ℝ) +implementationObjective t +implementationConstraints + (c : log (Vec.sum (Vec.exp ![x, y])) ≤ t) +solution (t := log ((exp x) + (exp y))) +solutionEqualsAtom by + rfl; +feasibility + (c : by + simp [Vec.sum, Vec.exp]) +optimality by + intros x' y' hx' hy' + simp [Vec.sum, Vec.exp] at c + refine le_trans ?_ c + rw [log_le_log (by positivity) (by positivity)] + apply add_le_add + { rw [exp_le_exp]; exact hx' } + { rw [exp_le_exp]; exact hy' } +vconditionElimination From 55fdf8dc90a72e2ae6b10cdba07c3001a0c4216f Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 11:23:51 -0500 Subject: [PATCH 085/189] feat: new atoms `min` and `max` --- CvxLean/Tactic/DCP/AtomLibrary/Fns/All.lean | 4 +++ CvxLean/Tactic/DCP/AtomLibrary/Fns/Max.lean | 28 +++++++++++++++++++++ CvxLean/Tactic/DCP/AtomLibrary/Fns/Min.lean | 28 +++++++++++++++++++++ 3 files changed, 60 insertions(+) create mode 100644 CvxLean/Tactic/DCP/AtomLibrary/Fns/Max.lean create mode 100644 CvxLean/Tactic/DCP/AtomLibrary/Fns/Min.lean diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/All.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/All.lean index 31fa3c72..49b4c896 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/All.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/All.lean @@ -12,8 +12,11 @@ import CvxLean.Tactic.DCP.AtomLibrary.Fns.Huber import CvxLean.Tactic.DCP.AtomLibrary.Fns.KLDiv import CvxLean.Tactic.DCP.AtomLibrary.Fns.Log import CvxLean.Tactic.DCP.AtomLibrary.Fns.LogDet +import CvxLean.Tactic.DCP.AtomLibrary.Fns.LogSumExp import CvxLean.Tactic.DCP.AtomLibrary.Fns.MatrixMul +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Max import CvxLean.Tactic.DCP.AtomLibrary.Fns.MaximizeNeg +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Min import CvxLean.Tactic.DCP.AtomLibrary.Fns.Mul import CvxLean.Tactic.DCP.AtomLibrary.Fns.MulVec import CvxLean.Tactic.DCP.AtomLibrary.Fns.Neg @@ -31,4 +34,5 @@ import CvxLean.Tactic.DCP.AtomLibrary.Fns.ToUpperTri import CvxLean.Tactic.DCP.AtomLibrary.Fns.Trace import CvxLean.Tactic.DCP.AtomLibrary.Fns.Transpose import CvxLean.Tactic.DCP.AtomLibrary.Fns.VecCons +import CvxLean.Tactic.DCP.AtomLibrary.Fns.VecConst import CvxLean.Tactic.DCP.AtomLibrary.Fns.XExp diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Max.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Max.lean new file mode 100644 index 00000000..089ed043 --- /dev/null +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Max.lean @@ -0,0 +1,28 @@ +import CvxLean.Tactic.DCP.Atoms +import CvxLean.Tactic.DCP.AtomLibrary.Sets.Le +import CvxLean.Lib.Math.Data.Vec + +namespace CvxLean + +open Real + +declare_atom max [convex] (x : ℝ)+ (y : ℝ)+ : max x y := +vconditions +implementationVars (t : ℝ) +implementationObjective t +implementationConstraints + (c1 : x ≤ t) + (c2 : y ≤ t) +solution (t := max x y) +solutionEqualsAtom by + rfl; +feasibility + (c1 : by simp) + (c2 : by simp) +optimality by + intros x' y' hx' hy' + rw [max_le_iff] + split_ands <;> linarith +vconditionElimination + +end CvxLean diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Min.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Min.lean new file mode 100644 index 00000000..8f7a04c6 --- /dev/null +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Min.lean @@ -0,0 +1,28 @@ +import CvxLean.Tactic.DCP.Atoms +import CvxLean.Tactic.DCP.AtomLibrary.Sets.Le +import CvxLean.Lib.Math.Data.Vec + +namespace CvxLean + +open Real + +declare_atom min [concave] (x : ℝ)+ (y : ℝ)+ : min x y := +vconditions +implementationVars (t : ℝ) +implementationObjective t +implementationConstraints + (c1 : t ≤ x) + (c2 : t ≤ y) +solution (t := min x y) +solutionEqualsAtom by + rfl; +feasibility + (c1 : by simp) + (c2 : by simp) +optimality by + intros x' y' hx' hy' + rw [le_min_iff] + split_ands <;> linarith +vconditionElimination + +end CvxLean From e24f08db158d11f4e4727f811fb22909057878e7 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 11:41:14 -0500 Subject: [PATCH 086/189] test: add quiz suite and short explanation for each suite --- rewriter/tests/test_almost_dgp.rs | 7 +++++-- rewriter/tests/test_cost_function.rs | 4 ++++ rewriter/tests/test_domain.rs | 6 +++++- rewriter/tests/test_dqcp.rs | 6 +++++- rewriter/tests/test_gp.rs | 4 ++++ rewriter/tests/test_main_example.rs | 6 +++++- rewriter/tests/test_misc.rs | 4 ++++ rewriter/tests/test_quiz.rs | 15 +++++++++++++++ rewriter/tests/test_rules.rs | 4 ++++ rewriter/tests/test_stanford.rs | 5 +++++ 10 files changed, 56 insertions(+), 5 deletions(-) create mode 100644 rewriter/tests/test_quiz.rs diff --git a/rewriter/tests/test_almost_dgp.rs b/rewriter/tests/test_almost_dgp.rs index f559e588..e6a33ad3 100644 --- a/rewriter/tests/test_almost_dgp.rs +++ b/rewriter/tests/test_almost_dgp.rs @@ -1,3 +1,7 @@ +/*! +Tests from geometric programming that do not follow the DGP rules. +!*/ + use egg_convexify::test_util::{*}; #[test] @@ -5,7 +9,6 @@ fn test_agp2() { convexify_check( "(exp (var x))", vec![ - "(le (mul (exp (var x)) (exp (var y))) (neg (div 2691 500)))" + "(le (sub (mul (exp (var x)) (exp (var y))) (div 2691 500)) 0)" ]); - } diff --git a/rewriter/tests/test_cost_function.rs b/rewriter/tests/test_cost_function.rs index acdefb69..b7478ec5 100644 --- a/rewriter/tests/test_cost_function.rs +++ b/rewriter/tests/test_cost_function.rs @@ -1,3 +1,7 @@ +/*! +Tests for cost function minimization. +!*/ + use egg_convexify::test_util::{*}; #[test] diff --git a/rewriter/tests/test_domain.rs b/rewriter/tests/test_domain.rs index 27809276..488877b6 100644 --- a/rewriter/tests/test_domain.rs +++ b/rewriter/tests/test_domain.rs @@ -1,10 +1,14 @@ +/*! +Tests for extended interval arithmetic. +!*/ + use egg_convexify; use egg_convexify::domain; use domain::Domain as Domain; -/* Tests for +, *, -, and / on critical intervals: +/* Tests for +, *, -, / and ^ on critical intervals: * Positive (0, +inf), * Non-negative [0, +inf), * Non-positive (-inf, 0], diff --git a/rewriter/tests/test_dqcp.rs b/rewriter/tests/test_dqcp.rs index ecdd4628..38f6894e 100644 --- a/rewriter/tests/test_dqcp.rs +++ b/rewriter/tests/test_dqcp.rs @@ -1,10 +1,14 @@ +/*! +Tests from quasiconvex programming. +!*/ + use egg_convexify::domain; use egg_convexify::test_util::{*}; #[test] -fn test_dqcp() { +fn test_qcp() { convexify_check_with_domain( vec![("x", domain::pos_dom())], "(var x)", diff --git a/rewriter/tests/test_gp.rs b/rewriter/tests/test_gp.rs index 48e05afa..bbc8979b 100644 --- a/rewriter/tests/test_gp.rs +++ b/rewriter/tests/test_gp.rs @@ -1,3 +1,7 @@ +/*! +Tests from geometric programming. +!*/ + use egg_convexify::domain; use egg_convexify::test_util::{*}; diff --git a/rewriter/tests/test_main_example.rs b/rewriter/tests/test_main_example.rs index 607906a3..ff391c31 100644 --- a/rewriter/tests/test_main_example.rs +++ b/rewriter/tests/test_main_example.rs @@ -1,3 +1,7 @@ +/*! +Test motivating example. +!*/ + use egg_convexify::domain; use egg_convexify::test_util::{*}; @@ -6,7 +10,7 @@ use egg_convexify::test_util::{*}; fn test_main_example() { convexify_check_with_domain( vec![("x", domain::pos_dom())], - "0", + "(var x)", vec![ "(le (div 1 (sqrt (var x))) (exp (var x)))" ]); diff --git a/rewriter/tests/test_misc.rs b/rewriter/tests/test_misc.rs index 1dc2aaba..e9c2cc5d 100644 --- a/rewriter/tests/test_misc.rs +++ b/rewriter/tests/test_misc.rs @@ -1,3 +1,7 @@ +/*! +Miscellaneous tests that do not fit in any other category. +!*/ + use egg_convexify::domain; use egg_convexify::test_util::{*}; diff --git a/rewriter/tests/test_quiz.rs b/rewriter/tests/test_quiz.rs new file mode 100644 index 00000000..0c4a1039 --- /dev/null +++ b/rewriter/tests/test_quiz.rs @@ -0,0 +1,15 @@ +/*! +Tests from random (non-DCP) problems generated by the DCP quiz: +https://dcp.stanford.edu/quiz +!*/ + +use egg_convexify::domain; + +use egg_convexify::test_util::{*}; + +#[test] +fn test_quiz_1() { + convexify_check_expression_with_domain( + vec![("x", domain::free_dom())], + "(inv (inv (var x)))"); +} diff --git a/rewriter/tests/test_rules.rs b/rewriter/tests/test_rules.rs index d859e0a6..73f5128e 100644 --- a/rewriter/tests/test_rules.rs +++ b/rewriter/tests/test_rules.rs @@ -1,3 +1,7 @@ +/*! +Tests for specific rewrite rules. +!*/ + use egg_convexify::domain; use egg_convexify::domain::Domain as Domain; diff --git a/rewriter/tests/test_stanford.rs b/rewriter/tests/test_stanford.rs index 48a974f1..7ea6f753 100644 --- a/rewriter/tests/test_stanford.rs +++ b/rewriter/tests/test_stanford.rs @@ -1,3 +1,8 @@ +/*! +Tests from the additional exercises in the Stanford Convex Optimization course: +https://github.com/cvxgrp/cvxbook_additional_exercises +!*/ + use egg_convexify::domain; use egg_convexify::test_util::{*}; From c53e8dc21de1276055d62dec86617edbc0acef3d Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 12:10:41 -0500 Subject: [PATCH 087/189] feat: align `pow` check with existing atoms --- rewriter/src/curvature.rs | 44 ++++++++++++++++++++++++++++++++++++--- 1 file changed, 41 insertions(+), 3 deletions(-) diff --git a/rewriter/src/curvature.rs b/rewriter/src/curvature.rs index 92f15786..617949f8 100644 --- a/rewriter/src/curvature.rs +++ b/rewriter/src/curvature.rs @@ -195,7 +195,6 @@ pub fn of_sub(c1: Curvature, c2: Curvature) -> Curvature { return of_add(c1, of_neg(c2)); } -// NOTE: we know that it is constant but we only have an interval. pub fn of_mul_by_const(c: Curvature, k: Domain) -> Curvature { if domain::is_neg(&k) { return of_neg(c); @@ -208,7 +207,8 @@ pub fn of_mul_by_const(c: Curvature, k: Domain) -> Curvature { } } -// NOTE: we know that it is constant but we only have an interval. +// This is a simplified version of `of_mul_by_const_full` that aligns with +// the existing power atoms. pub fn of_pow_by_const(c: Curvature, k: Domain, d_o: Option) -> Curvature { match c { Constant => { @@ -219,7 +219,45 @@ pub fn of_pow_by_const(c: Curvature, k: Domain, d_o: Option) -> Curvatur if domain::is_zero(&k) { return Constant; // Case x^1. - } else if k.subseteq(&Domain::make_cc(domain::one(), domain::one())) { + } else if domain::is_one(&k) { + return Affine; + // Case x^2. + } else if k.subseteq(&Domain::make_singleton(2.0)) { + return Convex; + // Case x^(-1). + } else if k.subseteq(&Domain::make_singleton(-1.0)) { + if domain::option_is_pos(d_o.as_ref()) { + return Convex; + } else { + return Unknown; + } + // Case x^(-2). + } else if k.subseteq(&Domain::make_singleton(-2.0)) { + if domain::option_is_pos(d_o.as_ref()) { + return Convex; + } else { + return Unknown; + } + } else { + return Unknown; + } + } + _ => { return Unknown; } + } +} + +#[allow(unused)] +pub fn of_pow_by_const_full(c: Curvature, k: Domain, d_o: Option) -> Curvature { + match c { + Constant => { + return Constant; + } + Affine => { + // Case x^0. + if domain::is_zero(&k) { + return Constant; + // Case x^1. + } else if domain::is_one(&k) { return Affine; // Case x^p with p < 0. } else if domain::is_neg(&k) { From fcd9f66dcbd9e8025e7239a8910e33777fdefb83 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 12:11:02 -0500 Subject: [PATCH 088/189] feat: simplify domain generation --- rewriter/src/domain.rs | 30 +++++++++++++++++------------- 1 file changed, 17 insertions(+), 13 deletions(-) diff --git a/rewriter/src/domain.rs b/rewriter/src/domain.rs index 9aa78171..4c026373 100644 --- a/rewriter/src/domain.rs +++ b/rewriter/src/domain.rs @@ -332,10 +332,8 @@ impl Serialize for Domain { /* Domain checks. */ -fn zero_ival() -> Interval { Interval::make(zero(), zero(), NO_ERROR) } - pub fn zero_dom() -> Domain { - Domain::make(zero_ival(), false, false) + Domain::make_cc(zero(), zero()) } pub fn is_zero(d: &Domain) -> bool { @@ -346,16 +344,24 @@ pub fn option_is_zero(d: Option<&Domain>) -> bool { d.map_or(false, is_zero) } -fn free_ival() -> Interval { Interval::make(neg_inf(), inf(), NO_ERROR) } +pub fn one_dom() -> Domain { + Domain::make_cc(one(), one()) +} -pub fn free_dom() -> Domain { - Domain::make(free_ival(), false, false) +pub fn is_one(d: &Domain) -> bool { + d.subseteq(&one_dom()) +} + +pub fn option_is_one(d: Option<&Domain>) -> bool { + d.map_or(false, is_one) } -fn nonneg_ival() -> Interval { Interval::make(zero(), inf(), NO_ERROR) } +pub fn free_dom() -> Domain { + Domain::make_ii() +} pub fn nonneg_dom() -> Domain { - Domain::make(nonneg_ival(), false, false) + Domain::make_ci(zero()) } pub fn is_nonneg(d: &Domain) -> bool { @@ -366,10 +372,8 @@ pub fn option_is_nonneg(d: Option<&Domain>) -> bool { d.map_or(false, is_nonneg) } -fn nonpos_ival() -> Interval { Interval::make(neg_inf(), neg_zero(), NO_ERROR) } - pub fn nonpos_dom() -> Domain { - Domain::make(nonpos_ival(), false, false) + Domain::make_ic(neg_zero()) } pub fn is_nonpos(d: &Domain) -> bool { @@ -381,7 +385,7 @@ pub fn option_is_nonpos(d: Option<&Domain>) -> bool { } pub fn pos_dom() -> Domain { - Domain::make(nonneg_ival(), true, false) + Domain::make_oi(zero()) } pub fn is_pos(d: &Domain) -> bool { @@ -393,7 +397,7 @@ pub fn option_is_pos(d: Option<&Domain>) -> bool { } pub fn neg_dom() -> Domain { - Domain::make(nonpos_ival(), false, true) + Domain::make_io(neg_zero()) } pub fn is_neg(d: &Domain) -> bool { From a90f4be8cfdf36fbfb4c7053a11da217b229b3e0 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 13:42:32 -0500 Subject: [PATCH 089/189] feat: domain arithmetic for `inv` and `abs` --- rewriter/src/domain.rs | 41 ++++++++++++++++++++++++++++++++++++----- 1 file changed, 36 insertions(+), 5 deletions(-) diff --git a/rewriter/src/domain.rs b/rewriter/src/domain.rs index 4c026373..fb71abfc 100644 --- a/rewriter/src/domain.rs +++ b/rewriter/src/domain.rs @@ -461,6 +461,34 @@ pub fn option_neg(d_o: Option) -> Option { execute_unary(d_o, neg) } +pub fn abs(d: &Domain) -> Domain { + let lo = d.lo_float().clone(); + let hi = d.hi_float().clone(); + let lo_open = d.lo_open; + let hi_open = d.hi_open; + let (lo_abs, hi_abs, lo_open_abs, hi_open_abs) = + if lo.is_sign_negative() { + if hi.is_sign_negative() { + (-hi, -lo, hi_open, lo_open) + } else { + if lo.abs() < hi.abs() { + (zero(), hi, false, hi_open) + } else if hi.abs() < lo.abs() { + (zero(), lo, false, lo_open) + } else { + (zero(), lo, false, lo_open && hi_open) + } + } + } else { + (lo, hi, lo_open, hi_open) + }; + Domain::make_from_endpoints(lo_abs, hi_abs, lo_open_abs, hi_open_abs) +} + +pub fn option_abs(d_o: Option) -> Option { + execute_unary(d_o, abs) +} + pub fn sqrt(d: &Domain) -> Domain { Domain::make(d.interval.sqrt(), d.lo_open, d.hi_open) } @@ -538,11 +566,6 @@ pub fn option_sub(d1_o: Option, d2_o: Option) -> Option execute_binary(d1_o, d2_o, sub) } -// The idea is that open "beats" closed. -fn choose_opennes(o_a: bool, o_b: bool) -> bool { - o_a || o_b -} - // Copied from interval-goods, but making multiplication by zero always be zero, // with the correct sign. fn perform_mult( @@ -730,6 +753,14 @@ pub fn option_div(d1_o: Option, d2_o: Option) -> Option execute_binary(d1_o, d2_o, div) } +pub fn inv(d: &Domain) -> Domain { + div(&one_dom(), d) +} + +pub fn option_inv(d_o: Option) -> Option { + execute_unary(d_o, inv) +} + // Same reasoning as in `perform_pow` fn perform_pow( lo1: &Float, lo2: &Float, From 7a1d64b86486e0c9392a49dda4775bc194093c7e Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 13:55:29 -0500 Subject: [PATCH 090/189] feat: domain arithmetic for `min` and `max` --- rewriter/src/domain.rs | 76 +++++++++++++++++++++++++++++++----------- 1 file changed, 57 insertions(+), 19 deletions(-) diff --git a/rewriter/src/domain.rs b/rewriter/src/domain.rs index fb71abfc..0777d238 100644 --- a/rewriter/src/domain.rs +++ b/rewriter/src/domain.rs @@ -462,27 +462,25 @@ pub fn option_neg(d_o: Option) -> Option { } pub fn abs(d: &Domain) -> Domain { - let lo = d.lo_float().clone(); - let hi = d.hi_float().clone(); - let lo_open = d.lo_open; - let hi_open = d.hi_open; - let (lo_abs, hi_abs, lo_open_abs, hi_open_abs) = - if lo.is_sign_negative() { - if hi.is_sign_negative() { - (-hi, -lo, hi_open, lo_open) + let a = d.lo_float().clone(); + let b = d.hi_float().clone(); + let l = d.lo_open; + let r = d.hi_open; + if a.is_sign_negative() { + if b.is_sign_negative() { + Domain::make_from_endpoints(-b, -a, r, l) + } else { + if a.abs() < b.abs() { + Domain::make_from_endpoints(zero(), b, false, r) + } else if b.abs() < a.abs() { + Domain::make_from_endpoints(zero(), a, false, l) } else { - if lo.abs() < hi.abs() { - (zero(), hi, false, hi_open) - } else if hi.abs() < lo.abs() { - (zero(), lo, false, lo_open) - } else { - (zero(), lo, false, lo_open && hi_open) - } + Domain::make_from_endpoints(zero(), a, false, l && r) } - } else { - (lo, hi, lo_open, hi_open) - }; - Domain::make_from_endpoints(lo_abs, hi_abs, lo_open_abs, hi_open_abs) + } + } else { + Domain::make_from_endpoints(a, b, l, r) + } } pub fn option_abs(d_o: Option) -> Option { @@ -534,6 +532,46 @@ pub fn option_entr(d_o: Option) -> Option { execute_unary(d_o, entr) } +pub fn min(d1: &Domain, d2: &Domain) -> Domain { + let a1 = d1.lo_float(); + let b1 = d1.hi_float(); + let a2 = d2.lo_float(); + let b2 = d2.hi_float(); + let l1 = d1.lo_open; + let r1 = d1.hi_open; + let l2 = d2.lo_open; + let r2 = d2.hi_open; + let lo = a1.min(a2); + let hi = b1.min(b2); + let lo_open = if a1 < a2 { l1 } else if a2 < a1 { l2 } else { l1 && l2 }; + let hi_open = if b1 < b2 { r1 } else if b2 < b1 { r2 } else { r1 || r2 }; + Domain::make_from_endpoints(lo, hi, lo_open, hi_open) +} + +pub fn option_min(d1_o: Option, d2_o: Option) -> Option { + execute_binary(d1_o, d2_o, min) +} + +pub fn max(d1: &Domain, d2: &Domain) -> Domain { + let a1 = d1.lo_float(); + let b1 = d1.hi_float(); + let a2 = d2.lo_float(); + let b2 = d2.hi_float(); + let l1 = d1.lo_open; + let r1 = d1.hi_open; + let l2 = d2.lo_open; + let r2 = d2.hi_open; + let lo = a1.max(a2); + let hi = b1.max(b2); + let lo_open = if a1 > a2 { l1 } else if a2 > a1 { l2 } else { l1 || l2 }; + let hi_open = if b1 > b2 { r1 } else if b2 > b1 { r2 } else { r1 && r2 }; + Domain::make_from_endpoints(lo, hi, lo_open, hi_open) +} + +pub fn option_max(d1_o: Option, d2_o: Option) -> Option { + execute_binary(d1_o, d2_o, max) +} + pub fn add(d1: &Domain, d2: &Domain) -> Domain { let l1 = d1.lo_open; let r1 = d1.hi_open; From 0fd5685d76e539dfbcb5d0300435cc2ffaa3a294 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 13:58:37 -0500 Subject: [PATCH 091/189] feat: domain arithmetic for `lse` --- rewriter/src/domain.rs | 14 +++++++++++--- 1 file changed, 11 insertions(+), 3 deletions(-) diff --git a/rewriter/src/domain.rs b/rewriter/src/domain.rs index 0777d238..8d1c938a 100644 --- a/rewriter/src/domain.rs +++ b/rewriter/src/domain.rs @@ -782,7 +782,7 @@ pub fn div(d1: &Domain, d2: &Domain) -> Domain { b1, b2, a1, b2, r1, r2, l1, r2) } else { - // Division by mixed (potentially zero), so the result is [-inf, inf]. + // Division by mixed (potentially zero), so the result is (-inf, inf). free_dom() } } @@ -799,7 +799,7 @@ pub fn option_inv(d_o: Option) -> Option { execute_unary(d_o, inv) } -// Same reasoning as in `perform_pow` +// Same reasoning as in `perform_mul`. fn perform_pow( lo1: &Float, lo2: &Float, hi1: &Float, hi2: &Float, @@ -836,7 +836,7 @@ pub fn pow(d1: &Domain, d2: &Domain) -> Domain { if !d1_nonneg { // We only define rules for nonnegative bases. // However, if the exponent is a constant, we consider some special - // cases: 0, 1, and even integers. + // cases: 0, 1, and even natural numbers. match d2.get_constant() { Some(f) => { if ((f as u32) as f64) == f { @@ -943,6 +943,14 @@ pub fn option_geo_mean(d1_o: Option, d2_o: Option) -> Option Domain { + log(&add(&exp(d1), &exp(d2))) +} + +pub fn option_log_sum_exp(d1_o: Option, d2_o: Option) -> Option { + execute_binary(d1_o, d2_o, log_sum_exp) +} + pub fn norm2(d1: &Domain, d2: &Domain) -> Domain { sqrt(&add(&pow(d1, &Domain::make_singleton(2.0)), &pow(d2, &Domain::make_singleton(2.0)))) } From 8bd7ce3309aa6a7e653fdcc4babd14304fedb26b Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 13:58:57 -0500 Subject: [PATCH 092/189] feat: extend language --- rewriter/src/optimization.rs | 28 ++++++++++++++++++++++++++-- 1 file changed, 26 insertions(+), 2 deletions(-) diff --git a/rewriter/src/optimization.rs b/rewriter/src/optimization.rs index b1844791..48a78601 100644 --- a/rewriter/src/optimization.rs +++ b/rewriter/src/optimization.rs @@ -16,11 +16,15 @@ define_language! { "eq" = Eq([Id; 2]), "le" = Le([Id; 2]), "neg" = Neg(Id), + "inv" = Inv(Id), + "abs" = Abs(Id), "sqrt" = Sqrt(Id), "log" = Log(Id), "exp" = Exp(Id), "xexp" = XExp(Id), "entr" = Entr(Id), + "min" = Min([Id; 2]), + "max" = Max([Id; 2]), "add" = Add([Id; 2]), "sub" = Sub([Id; 2]), "mul" = Mul([Id; 2]), @@ -28,6 +32,7 @@ define_language! { "pow" = Pow([Id; 2]), "qol" = QOL([Id; 2]), "geo" = Geo([Id; 2]), + "lse" = LSE([Id; 2]), "norm2" = Norm2([Id; 2]), "var" = Var(Id), "param" = Param(Id), @@ -110,6 +115,14 @@ impl Analysis for Meta { domain = domain::option_neg(get_domain(a)); is_constant = get_is_constant(a); } + Optimization::Inv(a) => { + domain = domain::option_inv(get_domain(a)); + is_constant = get_is_constant(a); + } + Optimization::Abs(a) => { + domain = domain::option_abs(get_domain(a)); + is_constant = get_is_constant(a); + } Optimization::Sqrt(a) => { domain = domain::option_sqrt(get_domain(a)); is_constant = get_is_constant(a); @@ -130,9 +143,16 @@ impl Analysis for Meta { domain = domain::option_entr(get_domain(a)); is_constant = get_is_constant(a); } + Optimization::Min([a, b]) => { + domain = domain::option_min(get_domain(a), get_domain(b)); + is_constant = get_is_constant(a) && get_is_constant(b); + } + Optimization::Max([a, b]) => { + domain = domain::option_max(get_domain(a), get_domain(b)); + is_constant = get_is_constant(a) && get_is_constant(b); + } Optimization::Add([a, b]) => { - domain = domain::option_add( - get_domain(a), get_domain(b)); + domain = domain::option_add(get_domain(a), get_domain(b)); is_constant = get_is_constant(a) && get_is_constant(b); } Optimization::Sub([a, b]) => { @@ -159,6 +179,10 @@ impl Analysis for Meta { domain = domain::option_geo_mean(get_domain(a), get_domain(b)); is_constant = get_is_constant(a) && get_is_constant(b); } + Optimization::LSE([a, b]) => { + domain = domain::option_log_sum_exp(get_domain(a), get_domain(b)); + is_constant = get_is_constant(a) && get_is_constant(b); + } Optimization::Norm2([a, b]) => { domain = domain::option_norm2(get_domain(a), get_domain(b)); is_constant = get_is_constant(a) && get_is_constant(b); From b2ad6aa75ba8ca2f50ce26f9c9e16fd723c5030c Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 14:19:08 -0500 Subject: [PATCH 093/189] feat: cost for new atoms and domain guards --- rewriter/src/cost.rs | 132 ++++++++++++++++++++++++++----------------- 1 file changed, 79 insertions(+), 53 deletions(-) diff --git a/rewriter/src/cost.rs b/rewriter/src/cost.rs index f3f100b0..6e1880ae 100644 --- a/rewriter/src/cost.rs +++ b/rewriter/src/cost.rs @@ -91,12 +91,9 @@ impl<'a> CostFunction for DCPCost<'a> { } } Optimization::Eq([a, b]) => { - curvature = - if get_curvature!(a) <= Curvature::Affine && get_curvature!(b) <= Curvature::Affine { - Curvature::Affine - } else { - Curvature::Unknown - }; + if get_curvature!(a) <= Curvature::Affine && get_curvature!(b) <= Curvature::Affine { + curvature = Curvature::Affine + } num_vars = get_num_vars!(a) + get_num_vars!(b); term_size = 1 + get_term_size!(a) + get_term_size!(b); } @@ -110,13 +107,41 @@ impl<'a> CostFunction for DCPCost<'a> { num_vars = get_num_vars!(a); term_size = 1 + get_term_size!(a); } + Optimization::Inv(a) => { + let da_pos = domain::option_is_pos(get_domain(a).as_ref()); + if da_pos { + curvature = curvature::of_convex_nonincreasing_fn(get_curvature!(a)); + } + num_vars = get_num_vars!(a); + term_size = 1 + get_term_size!(a); + } + Optimization::Abs(a) => { + let da_nonneg = domain::option_is_nonneg(get_domain(a).as_ref()); + let da_nonpos = domain::option_is_nonpos(get_domain(a).as_ref()); + if da_nonneg { + curvature = curvature::of_convex_nondecreasing_fn(get_curvature!(a)); + } else if da_nonpos { + curvature = curvature::of_convex_nonincreasing_fn(get_curvature!(a)); + } else { + curvature = curvature::of_convex_none_fn(get_curvature!(a)); + } + num_vars = get_num_vars!(a); + term_size = 1 + get_term_size!(a); + } Optimization::Sqrt(a) => { + let da_nonneg = domain::option_is_nonneg(get_domain(a).as_ref()); + if da_nonneg { + curvature = curvature::of_concave_nondecreasing_fn(get_curvature!(a)); + } curvature = curvature::of_concave_nondecreasing_fn(get_curvature!(a)); num_vars = get_num_vars!(a); term_size = 1 + get_term_size!(a); } Optimization::Log(a) => { - curvature = curvature::of_concave_nondecreasing_fn(get_curvature!(a)); + let da_nonneg = domain::option_is_nonneg(get_domain(a).as_ref()); + if da_nonneg { + curvature = curvature::of_concave_nondecreasing_fn(get_curvature!(a)); + } num_vars = get_num_vars!(a); term_size = 1 + get_term_size!(a); } @@ -126,12 +151,18 @@ impl<'a> CostFunction for DCPCost<'a> { term_size = 1 + get_term_size!(a); } Optimization::XExp(a) => { - curvature = curvature::of_convex_nondecreasing_fn(get_curvature!(a)); + let da_nonneg = domain::option_is_nonneg(get_domain(a).as_ref()); + if da_nonneg { + curvature = curvature::of_convex_nondecreasing_fn(get_curvature!(a)); + } num_vars = get_num_vars!(a); term_size = 1 + get_term_size!(a); } Optimization::Entr(a) => { - curvature = curvature::of_concave_none_fn(get_curvature!(a)); + let da_pos = domain::option_is_pos(get_domain(a).as_ref()); + if da_pos { + curvature = curvature::of_concave_none_fn(get_curvature!(a)); + } num_vars = get_num_vars!(a); term_size = 1 + get_term_size!(a); } @@ -148,73 +179,54 @@ impl<'a> CostFunction for DCPCost<'a> { Optimization::Mul([a, b]) => { let da_o = get_domain(a); let db_o = get_domain(b); - curvature = match (get_is_constant(a), get_is_constant(b)) { + match (get_is_constant(a), get_is_constant(b)) { (true, true) => { - Curvature::Constant + curvature = Curvature::Constant } (true, false) => { - match da_o { - Some(da) => { - curvature::of_mul_by_const(get_curvature!(b), da) - } - None => { Curvature::Unknown } + if let Some(da) = da_o { + curvature = curvature::of_mul_by_const(get_curvature!(b), da) } } (false, true) => { - match db_o { - Some(db) => { - curvature::of_mul_by_const(get_curvature!(a), db) - } - None => { Curvature::Unknown } + if let Some(db) = db_o { + curvature = curvature::of_mul_by_const(get_curvature!(a), db) } } - _ => { Curvature::Unknown } - }; + _ => { } + } num_vars = get_num_vars!(a) + get_num_vars!(b); term_size = 1 + get_term_size!(a) + get_term_size!(b); } Optimization::Div([a, b]) => { let db_o = get_domain(b); - curvature = match (get_is_constant(a), get_is_constant(b)) { + match (get_is_constant(a), get_is_constant(b)) { (true, true) => { - match db_o { - Some(db) => { - if domain::does_not_contain_zero(&db) { - Curvature::Constant - } else { - Curvature::Unknown - } - + if let Some(db) = db_o { + if domain::does_not_contain_zero(&db) { + curvature = Curvature::Constant } - None => { Curvature::Unknown } } } (false, true) => { - match db_o { - Some(db) => { - if domain::does_not_contain_zero(&db) { - curvature::of_mul_by_const(get_curvature!(a), db) - } else { - Curvature::Unknown - } + if let Some(db) = db_o { + if domain::does_not_contain_zero(&db) { + curvature = curvature::of_mul_by_const(get_curvature!(a), db) } - None => { Curvature::Unknown } } } - _ => { Curvature::Unknown } + _ => { } }; num_vars = get_num_vars!(a) + get_num_vars!(b); term_size = 1 + get_term_size!(a) + get_term_size!(b); } Optimization::Pow([a, b]) => { - curvature = if get_is_constant(b) { - match get_domain(b) { - Some(db) => { - curvature::of_pow_by_const(get_curvature!(a), db, get_domain(a)) - } - _ => { Curvature::Unknown } + if get_is_constant(b) { + if let Some(db) = get_domain(b) { + // Domain guards already in `of_pow_by_const`. + curvature = curvature::of_pow_by_const(get_curvature!(a), db, get_domain(a)) } - } else { Curvature::Unknown }; + } num_vars = get_num_vars!(a) + get_num_vars!(b); term_size = 1 + get_term_size!(a) + get_term_size!(b); } @@ -229,14 +241,28 @@ impl<'a> CostFunction for DCPCost<'a> { } else { curvature::of_convex_none_fn(get_curvature!(a)) }; - let curvature_den = curvature::of_convex_nonincreasing_fn(get_curvature!(b)); - curvature = curvature::join(curvature_num, curvature_den); + let db_pos = domain::option_is_pos(get_domain(b).as_ref()); + if db_pos { + let curvature_den = curvature::of_convex_nonincreasing_fn(get_curvature!(b)); + curvature = curvature::join(curvature_num, curvature_den); + } num_vars = get_num_vars!(a) + get_num_vars!(b); term_size = 1 + get_term_size!(a) + get_term_size!(b); } Optimization::Geo([a, b]) => { - let curvature_a = curvature::of_concave_nondecreasing_fn(get_curvature!(a)); - let curvature_b = curvature::of_concave_nondecreasing_fn(get_curvature!(b)); + let da_nonneg = domain::option_is_nonneg(get_domain(a).as_ref()); + let db_nonneg = domain::option_is_nonneg(get_domain(b).as_ref()); + if da_nonneg && db_nonneg { + let curvature_a = curvature::of_concave_nondecreasing_fn(get_curvature!(a)); + let curvature_b = curvature::of_concave_nondecreasing_fn(get_curvature!(b)); + curvature = curvature::join(curvature_a, curvature_b); + } + num_vars = get_num_vars!(a) + get_num_vars!(b); + term_size = 1 + get_term_size!(a) + get_term_size!(b); + } + Optimization::LSE([a, b]) => { + let curvature_a = curvature::of_convex_nondecreasing_fn(get_curvature!(a)); + let curvature_b = curvature::of_convex_nondecreasing_fn(get_curvature!(b)); curvature = curvature::join(curvature_a, curvature_b); num_vars = get_num_vars!(a) + get_num_vars!(b); term_size = 1 + get_term_size!(a) + get_term_size!(b); From 6fd8ad4ce30d226530bdbafbcd02f9c6008cb46a Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 14:21:14 -0500 Subject: [PATCH 094/189] feat: cost for `min` and `max` --- rewriter/src/cost.rs | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/rewriter/src/cost.rs b/rewriter/src/cost.rs index 6e1880ae..d969fa87 100644 --- a/rewriter/src/cost.rs +++ b/rewriter/src/cost.rs @@ -166,6 +166,20 @@ impl<'a> CostFunction for DCPCost<'a> { num_vars = get_num_vars!(a); term_size = 1 + get_term_size!(a); } + Optimization::Min([a, b]) => { + let curvature_a = curvature::of_concave_nondecreasing_fn(get_curvature!(a)); + let curvature_b = curvature::of_concave_nondecreasing_fn(get_curvature!(b)); + curvature = curvature::join(curvature_a, curvature_b); + num_vars = get_num_vars!(a) + get_num_vars!(b); + term_size = 1 + get_term_size!(a) + get_term_size!(b); + } + Optimization::Max([a, b]) => { + let curvature_a = curvature::of_convex_nondecreasing_fn(get_curvature!(a)); + let curvature_b = curvature::of_convex_nondecreasing_fn(get_curvature!(b)); + curvature = curvature::join(curvature_a, curvature_b); + num_vars = get_num_vars!(a) + get_num_vars!(b); + term_size = 1 + get_term_size!(a) + get_term_size!(b); + } Optimization::Add([a, b]) => { curvature = curvature::of_add(get_curvature!(a), get_curvature!(b)); term_size = 1 + get_term_size!(a) + get_term_size!(b); From a948c11ca51174766f29e4746e2c70078be39007 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 14:24:35 -0500 Subject: [PATCH 095/189] fix: borrowing in `min` and `max` --- rewriter/src/domain.rs | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/rewriter/src/domain.rs b/rewriter/src/domain.rs index 8d1c938a..ba4236da 100644 --- a/rewriter/src/domain.rs +++ b/rewriter/src/domain.rs @@ -541,8 +541,8 @@ pub fn min(d1: &Domain, d2: &Domain) -> Domain { let r1 = d1.hi_open; let l2 = d2.lo_open; let r2 = d2.hi_open; - let lo = a1.min(a2); - let hi = b1.min(b2); + let lo = a1.clone().min(a2); + let hi = b1.clone().min(b2); let lo_open = if a1 < a2 { l1 } else if a2 < a1 { l2 } else { l1 && l2 }; let hi_open = if b1 < b2 { r1 } else if b2 < b1 { r2 } else { r1 || r2 }; Domain::make_from_endpoints(lo, hi, lo_open, hi_open) @@ -561,8 +561,8 @@ pub fn max(d1: &Domain, d2: &Domain) -> Domain { let r1 = d1.hi_open; let l2 = d2.lo_open; let r2 = d2.hi_open; - let lo = a1.max(a2); - let hi = b1.max(b2); + let lo = a1.clone().max(a2); + let hi = b1.clone().max(b2); let lo_open = if a1 > a2 { l1 } else if a2 > a1 { l2 } else { l1 || l2 }; let hi_open = if b1 > b2 { r1 } else if b2 > b1 { r2 } else { r1 && r2 }; Domain::make_from_endpoints(lo, hi, lo_open, hi_open) From 24035d05dccf17d1442cd29ab1eccfce0eae815c Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 14:25:15 -0500 Subject: [PATCH 096/189] doc: small comment about DCP in cost calculation --- rewriter/src/cost.rs | 1 + 1 file changed, 1 insertion(+) diff --git a/rewriter/src/cost.rs b/rewriter/src/cost.rs index d969fa87..96793a4f 100644 --- a/rewriter/src/cost.rs +++ b/rewriter/src/cost.rs @@ -17,6 +17,7 @@ impl<'a> CostFunction for DCPCost<'a> { // Curvature + number of variables (with repetition) + term size. // In lexicographic order. type Cost = (Curvature, u32, u32); + // Curvature analysis corresponds exactly to the DCP rules. fn cost(&mut self, enode: &Optimization, mut costs: C) -> Self::Cost where C: FnMut(Id) -> Self::Cost From 97f65dfbb922f3c85f7a6c1f2e2b4421c75ab0ca Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 14:27:43 -0500 Subject: [PATCH 097/189] fix: borrowing in `abs` --- rewriter/src/domain.rs | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/rewriter/src/domain.rs b/rewriter/src/domain.rs index ba4236da..58632797 100644 --- a/rewriter/src/domain.rs +++ b/rewriter/src/domain.rs @@ -463,16 +463,18 @@ pub fn option_neg(d_o: Option) -> Option { pub fn abs(d: &Domain) -> Domain { let a = d.lo_float().clone(); + let a_abs = a.clone().abs(); let b = d.hi_float().clone(); + let b_abs = b.clone().abs(); let l = d.lo_open; let r = d.hi_open; if a.is_sign_negative() { if b.is_sign_negative() { Domain::make_from_endpoints(-b, -a, r, l) } else { - if a.abs() < b.abs() { + if a_abs < b_abs { Domain::make_from_endpoints(zero(), b, false, r) - } else if b.abs() < a.abs() { + } else if b_abs < a_abs { Domain::make_from_endpoints(zero(), a, false, l) } else { Domain::make_from_endpoints(zero(), a, false, l && r) From e458a1885ebc708d4f915e92f10c6fd8a17af5e8 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 14:28:37 -0500 Subject: [PATCH 098/189] fix: curvature in `sqrt` only for nonnegative arguments --- rewriter/src/cost.rs | 1 - 1 file changed, 1 deletion(-) diff --git a/rewriter/src/cost.rs b/rewriter/src/cost.rs index 96793a4f..d9e64fba 100644 --- a/rewriter/src/cost.rs +++ b/rewriter/src/cost.rs @@ -134,7 +134,6 @@ impl<'a> CostFunction for DCPCost<'a> { if da_nonneg { curvature = curvature::of_concave_nondecreasing_fn(get_curvature!(a)); } - curvature = curvature::of_concave_nondecreasing_fn(get_curvature!(a)); num_vars = get_num_vars!(a); term_size = 1 + get_term_size!(a); } From cf3a8ce98bc3c364b4ac8330106dace41a9368d3 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 16:18:18 -0500 Subject: [PATCH 099/189] feat: extend `prepare_positivity` --- CvxLean/Tactic/Util/PositivityExt.lean | 54 +++++++++++++++++++------- 1 file changed, 40 insertions(+), 14 deletions(-) diff --git a/CvxLean/Tactic/Util/PositivityExt.lean b/CvxLean/Tactic/Util/PositivityExt.lean index 07f5a7f7..5bb97534 100644 --- a/CvxLean/Tactic/Util/PositivityExt.lean +++ b/CvxLean/Tactic/Util/PositivityExt.lean @@ -14,28 +14,54 @@ elab (name := cases_and) "cases_and" : tactic => do def preparePositivity (mvarId : MVarId) : MetaM MVarId := do mvarId.withContext do + -- Adjust hypotheses if needed. let mut hyps := #[] + let le_lemmas := [``sub_nonneg_of_le, ``neg_nonneg_of_nonpos] + let lt_lemmas := [``sub_pos_of_lt, ``neg_pos_of_neg] let mut lctx ← getLCtx for localDecl in lctx do let ty := localDecl.type - let le_lemma := ``sub_nonneg_of_le - let lt_lemma := ``sub_pos_of_lt - for (le_or_lt, le_or_lt_lemma) in [(``LE.le, le_lemma), (``LT.lt, lt_lemma)] do + for (le_or_lt, lemmas) in [(``LE.le, le_lemmas), (``LT.lt, lt_lemmas)] do let ty := Expr.consumeMData ty match ty.app4? le_or_lt with - | some (R, _, lhs, _rhs) => - if ← isDefEq R q(ℝ) then - if ← isDefEq lhs q(0 : ℝ) then - continue - -- If LHS is not zero, add new hypothesis. - let val ← mkAppM le_or_lt_lemma #[localDecl.toExpr] - let ty ← inferType val - let n := localDecl.userName - hyps := hyps.push (Hypothesis.mk n ty val) + | some (R, _, lhs, rhs) => + if !(← isDefEq R q(ℝ)) then + continue + if ← isDefEq lhs q(0 : ℝ) then + continue + let le_or_lt_lemma := + if ← isDefEq rhs q(0 : ℝ) then lemmas[1]! else lemmas[0]! + -- If LHS is not zero, add new hypothesis. + let val ← mkAppM le_or_lt_lemma #[localDecl.toExpr] + let ty ← inferType val + let n := localDecl.userName + hyps := hyps.push (Hypothesis.mk n ty val) | none => continue - let (_, newMVarId) ← mvarId.assertHypotheses hyps - return newMVarId + let (_, mvarId) ← mvarId.assertHypotheses hyps + + -- Adjust goal if needed. + let goalExpr ← mvarId.getType + let mut mvarId := mvarId + let le_lemmas := [``le_of_sub_nonneg, ``nonpos_of_neg_nonneg] + let lt_lemmas := [``lt_of_sub_pos, ``neg_of_neg_pos] + for (le_or_lt, lemmas) in [(``LE.le, le_lemmas), (``LT.lt, lt_lemmas)] do + match goalExpr.app4? le_or_lt with + | some (R, _, lhs, rhs) => + if !(← isDefEq R q(ℝ)) then + continue + if ← isDefEq lhs q(0 : ℝ) then + continue + let le_or_lt_lemma := + if ← isDefEq rhs q(0 : ℝ) then lemmas[1]! else lemmas[0]! + if let [g] ← mvarId.applyConst le_or_lt_lemma then + mvarId := g + break + else + throwError "prepare_positivity failed" + | none => continue + + return mvarId elab (name := prepare_positivity) "prepare_positivity" : tactic => do let mvarId ← getMainGoal From 562d1a4d7cbdc2fb2d43bf8fa2f8913958ba7bbf Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 16:28:26 -0500 Subject: [PATCH 100/189] refactor: `positivity_ext` -> `arith` --- .../PositivityExt.lean => Arith/Arith.lean} | 4 +- CvxLean/Tactic/Arith/PositvityExt.lean | 108 ++++++++++++++++++ .../Tactic/Convexify/RewriteMapLibrary.lean | 60 +++++----- CvxLean/Tactic/DCP/Dcp.lean | 2 +- CvxLean/Test/Convexify/DQCP.lean | 76 ------------ CvxLean/Test/Convexify/Unit.lean | 50 -------- 6 files changed, 141 insertions(+), 159 deletions(-) rename CvxLean/Tactic/{Util/PositivityExt.lean => Arith/Arith.lean} (97%) create mode 100644 CvxLean/Tactic/Arith/PositvityExt.lean diff --git a/CvxLean/Tactic/Util/PositivityExt.lean b/CvxLean/Tactic/Arith/Arith.lean similarity index 97% rename from CvxLean/Tactic/Util/PositivityExt.lean rename to CvxLean/Tactic/Arith/Arith.lean index 5bb97534..e87a7913 100644 --- a/CvxLean/Tactic/Util/PositivityExt.lean +++ b/CvxLean/Tactic/Arith/Arith.lean @@ -70,8 +70,8 @@ elab (name := prepare_positivity) "prepare_positivity" : tactic => do end Tactic -syntax "positivity_ext" : tactic +syntax "arith" : tactic macro_rules - | `(tactic| positivity_ext) => + | `(tactic| arith) => `(tactic| cases_and; prepare_positivity; positivity) diff --git a/CvxLean/Tactic/Arith/PositvityExt.lean b/CvxLean/Tactic/Arith/PositvityExt.lean new file mode 100644 index 00000000..786443ca --- /dev/null +++ b/CvxLean/Tactic/Arith/PositvityExt.lean @@ -0,0 +1,108 @@ +import Lean +import Qq +import Mathlib.Data.Real.Basic +import Mathlib.Tactic.Positivity +import CvxLean.Lib.Math.Data.Real + +namespace Mathlib.Meta.Positivity + +open Lean.Meta Qq + +lemma Real.one_sub_one_div_sq_nonneg_of_le_one {x : ℝ} : + 0 < x → 0 ≤ 1 - x → 0 ≤ (1 / x ^ 2) - 1 := + fun h1 h2 => by + have hx1 : x ≤ 1 := by linarith + have h0x : 0 ≤ x := by positivity + have h0x2 : 0 < x ^ 2 := by positivity + rw [le_sub_iff_add_le, zero_add, le_div_iff h0x2, one_mul] + simpa [abs_eq_self.mpr h0x] + +@[positivity ((1 / (_ : ℝ) ^ (2 : ℝ)) - 1)] +def evalOneDivSqSubOne : PositivityExt where eval {_ _α} zα pα e := do + let (.app (.app _sub (.app (.app _div _one) (.app (.app _pow (x : Q(ℝ))) _two))) _one') ← + withReducible (whnf e) | throwError "not ((1 / x ^ 2) - 1)" + -- 0 < x ? + let h1 := + match ← core zα pα x with + | .positive pa => some pa + | _ => none + -- 0 ≤ 1 - x ? + let h2 := ← do + match ← core zα pα (q((1 : ℝ) - $x) : Q(ℝ)) with + | .nonnegative pa => return some pa + | .positive pa => + let pa ← mkAppM ``le_of_lt #[pa] + return some pa + | _ => return none + -- If 0 < x and 0 ≤ 1 - x, then 0 ≤ (1 / x ^ 2) - 1 + match h1, h2 with + | some h1', some h2' => + let pa' ← mkAppM ``Real.one_sub_one_div_sq_nonneg_of_le_one #[h1', h2'] + pure (.nonnegative pa') + | _, _ => + pure .none + +lemma Real.one_sub_sq_nonneg_of_le_one {x : ℝ} : + 0 ≤ x → 0 ≤ 1 - x → 0 ≤ 1 - x ^ (2 : ℝ) := + fun h1 h2 => by + have hx1 : x ≤ 1 := by linarith + rw [le_sub_iff_add_le, zero_add] + simpa [abs_eq_self.mpr h1] + +@[positivity (1 - ((_ : ℝ) ^ (2 : ℝ)))] +def evalOneSubSq : PositivityExt where eval {_ _α} zα pα e := do + let (.app (.app _sub _one) (.app (.app _pow (x : Q(ℝ))) _two)) ← + withReducible (whnf e) | throwError "not (1 - x ^ 2)" + -- 0 ≤ x ? + let h1 := ← do + match ← core zα pα x with + | .nonnegative pa => return some pa + | .positive pa => + let pa ← mkAppM ``le_of_lt #[pa] + return some pa + | _ => return none + -- 0 ≤ 1 - x ? + let h2 := ← do + match ← core zα pα (q((1 : ℝ) - $x) : Q(ℝ)) with + | .nonnegative pa => return some pa + | .positive pa => + let pa ← mkAppM ``le_of_lt #[pa] + return some pa + | _ => return none + -- If 0 ≤ x and 0 ≤ 1 - x, then 0 ≤ 1 - x ^ 2 + match h1, h2 with + | some h1', some h2' => + let pa' ← mkAppM ``Real.one_sub_sq_nonneg_of_le_one #[h1', h2'] + pure (.nonnegative pa') + | _, _ => + pure .none + +lemma Real.exp_sub_one_pos_of_pos {x : ℝ} : 0 < x → 0 < Real.exp x - 1 := + fun h => by simp [h] + +@[positivity ((Real.exp (_ : ℝ)) - 1)] +def evalExpSubOne : PositivityExt where eval {_ _α} zα pα e := do + let (.app (.app _sub (.app _exp (x : Q(ℝ)))) _one) ← + withReducible (whnf e) | throwError "not (Real.exp x - 1)" + match ← core zα pα x with + | .positive pa => + let pa' ← mkAppM ``Real.exp_sub_one_pos_of_pos #[pa] + pure (.positive pa') + | _ => + pure .none + +lemma Real.one_sub_div_exp_pos_of_pos {x : ℝ} : 0 < x → 0 < 1 - 1 / Real.exp x := + fun h => by field_simp; positivity + +@[positivity (1 - (1 / (Real.exp (_ : ℝ))))] +def evalOneSubDivExp : PositivityExt where eval {_ _α} zα pα e := do + let (.app (.app _sub _one) (.app (.app _div _one') (.app _exp (x : Q(ℝ))))) ← + withReducible (whnf e) | throwError "not (1 - 1 / Real.exp x)" + match ← core zα pα x with + | .positive pa => + let pa' ← mkAppM ``Real.one_sub_div_exp_pos_of_pos #[pa] + pure (.positive pa') + | _ => + pure .none + +end Mathlib.Meta.Positivity diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index a3f4a01a..ad7ca14a 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -1,6 +1,6 @@ import CvxLean.Tactic.Convexify.RewriteMapCmd import CvxLean.Tactic.Convexify.Basic -import CvxLean.Tactic.Util.PositivityExt +import CvxLean.Tactic.Arith.Basic -- TODO: Move. lemma Real.log_eq_log {x y : ℝ} (hx : 0 < x) (hy : 0 < y) : Real.log x = Real.log y ↔ x = y := @@ -65,7 +65,7 @@ register_objFun_rewrite_map "map_objFun_sq"; "(prob (objFun ?a) ?cs)" => "(prob /- Equality rules. -/ register_rewrite_map "log_eq_log" ; "(eq ?a ?b)" => "(eq (log ?a) (log ?b))" := - rewrite [Real.log_eq_log (by positivity_ext) (by positivity_ext)]; + rewrite [Real.log_eq_log (by arith) (by arith)]; /- Less than or equal rules. -/ @@ -83,25 +83,25 @@ register_rewrite_map "sub_le_iff_le_add-rev"; "(le ?a (add ?b ?c))" => "(le (sub rewrite [←sub_le_iff_le_add]; register_rewrite_map "div_le_iff" ; "(le (div ?a ?b) ?c)" => "(le ?a (mul ?b ?c))" := - rewrite [div_le_iff (by positivity_ext)]; + rewrite [div_le_iff (by arith)]; register_rewrite_map "div_le_iff-rev" ; "(le (div ?a ?b) ?c)" => "(le ?a (mul ?b ?c))" := - rewrite [←div_le_iff (by positivity_ext)]; + rewrite [←div_le_iff (by arith)]; register_rewrite_map "div_le_one-rev" ; "(le ?a ?b)" => "(le (div ?a ?b) 1)" := - rewrite [←div_le_one (by positivity_ext)]; + rewrite [←div_le_one (by arith)]; register_rewrite_map "log_le_log" ; "(le (log ?a) (log ?b))" => "(le ?a ?b)" := - rewrite [Real.log_le_log (by positivity_ext) (by positivity_ext)]; + rewrite [Real.log_le_log (by arith) (by arith)]; register_rewrite_map "log_le_log-rev" ; "(le ?a ?b)" => "(le (log ?a) (log ?b))" := - rewrite [←Real.log_le_log (by positivity_ext) (by positivity_ext)]; + rewrite [←Real.log_le_log (by arith) (by arith)]; register_rewrite_map "pow_two_le_pow_two"; "(le (pow ?a 2) (pow ?b 2))" => "(le ?a ?b)":= - rewrite [Real.pow_two_le_pow_two (by positivity_ext) (by positivity_ext)]; + rewrite [Real.pow_two_le_pow_two (by arith) (by arith)]; register_rewrite_map "pow_two_le_pow_two-rev"; "(le ?a ?b)" => "(le (pow ?a 2) (pow ?b 2))" := - rewrite [←Real.pow_two_le_pow_two (by positivity_ext) (by positivity_ext)]; + rewrite [←Real.pow_two_le_pow_two (by arith) (by arith)]; /- Field rules. -/ @@ -180,7 +180,7 @@ register_rewrite_map "mul_div-rev" ; "(div (mul ?a ?b) ?c)" => "(mul ?a (div ?b simp_or_rw [←mul_div (G := ℝ)]; register_rewrite_map "div_self" ; "(div ?a ?a)" => "1" := - simp_or_rw [div_self (G₀ := ℝ) (by positivity_ext)]; + simp_or_rw [div_self (G₀ := ℝ) (by arith)]; /- Power and square root rules. -/ @@ -192,43 +192,43 @@ register_rewrite_map "pow_one"; "(pow ?a 1)" => "?a" := simp_or_rw [Real.rpow_one]; register_rewrite_map "pow_add"; "(pow ?a (add ?b ?c))" => "(mul (pow ?a ?b) (pow ?a ?c))" := - simp_or_rw [Real.rpow_add (by positivity_ext)]; + simp_or_rw [Real.rpow_add (by arith)]; register_rewrite_map "pow_add-rev"; "(mul (pow ?a ?b) (pow ?a ?c))" => "(pow ?a (add ?b ?c))" := - simp_or_rw [←Real.rpow_add (by positivity_ext)]; + simp_or_rw [←Real.rpow_add (by arith)]; register_rewrite_map "mul_pow"; "(mul (pow ?a ?n) (pow ?b ?n))" => "(pow (mul ?a ?b) ?n)" := - simp_or_rw [Real.mul_rpow (by positivity_ext) (by positivity_ext)]; + simp_or_rw [Real.mul_rpow (by arith) (by arith)]; register_rewrite_map "mul_pow-rev"; "(mul (pow ?a ?n) (pow ?b ?n))" => "(pow (mul ?a ?b) ?n)" := - simp_or_rw [←Real.mul_rpow (by positivity_ext) (by positivity_ext)]; + simp_or_rw [←Real.mul_rpow (by arith) (by arith)]; register_rewrite_map "pow_mul"; "(pow ?a (mul ?n ?m))" => "(pow (pow ?a ?n) ?m)" := - simp_or_rw [Real.rpow_mul (by positivity_ext)]; + simp_or_rw [Real.rpow_mul (by arith)]; register_rewrite_map "pow_mul-rev"; "(pow (pow ?a ?n) ?m)" => "(pow ?a (mul ?n ?m))" := - simp_or_rw [←Real.rpow_mul (by positivity_ext)]; + simp_or_rw [←Real.rpow_mul (by arith)]; register_rewrite_map "div_pow"; "(pow (div ?a ?b) ?n)" => "(div (pow ?a ?n) (pow ?b ?n))" := - simp_or_rw [Real.div_rpow (by positivity_ext) (by positivity_ext)]; + simp_or_rw [Real.div_rpow (by arith) (by arith)]; register_rewrite_map "div_pow-rev"; "(div (pow ?a ?n) (pow ?b ?n))" => "(pow (div ?a ?b) ?n)" := - simp_or_rw [←Real.div_rpow (by positivity_ext) (by positivity_ext)]; + simp_or_rw [←Real.div_rpow (by arith) (by arith)]; register_rewrite_map "pow_sub"; "(pow ?a (sub ?b ?c))" => "(div (pow ?a ?b) (pow ?a ?c))" := - simp_or_rw [Real.rpow_sub (by positivity_ext)]; + simp_or_rw [Real.rpow_sub (by arith)]; register_rewrite_map "pow_sub-rev"; "(div (pow ?a ?b) (pow ?a ?c))" => "(pow ?a (sub ?b ?c))" := - simp_or_rw [←Real.rpow_sub (by positivity_ext)]; + simp_or_rw [←Real.rpow_sub (by arith)]; register_rewrite_map "div_pow_eq_mul_pow_neg" ; "(div ?a (pow ?b ?c))" => "(mul ?a (pow ?b (neg ?c)))" := - simp_or_rw [Real.div_pow_eq_mul_pow_neg (by positivity_ext)]; + simp_or_rw [Real.div_pow_eq_mul_pow_neg (by arith)]; register_rewrite_map "div_pow_eq_mul_pow_neg-rev" ; "(div ?a (pow ?b ?c))" => "(mul ?a (pow ?b (neg ?c)))" := - simp_or_rw [←Real.div_pow_eq_mul_pow_neg (by positivity_ext)]; + simp_or_rw [←Real.div_pow_eq_mul_pow_neg (by arith)]; register_rewrite_map "one_div_eq_pow_neg_one"; "(div 1 ?a)" => "(pow ?a (neg 1))" := - simp_or_rw [Real.one_div_eq_pow_neg_one (by positivity_ext)]; + simp_or_rw [Real.one_div_eq_pow_neg_one (by arith)]; register_rewrite_map "sqrt_eq_rpow" ; "(sqrt ?a)" => "(pow ?a 0.5)" := simp_or_rw [Real.sqrt_eq_rpow]; @@ -243,11 +243,11 @@ register_rewrite_map "pow_two-rev"; "(mul ?a ?a)" => "(pow ?a 2)" := simp_or_rw [←pow_two (M := ℝ)]; register_rewrite_map "pow_half_two"; "(pow (pow ?a 0.5) 2)" => "?a" := - simp_or_rw [Real.pow_half_two (by positivity_ext)]; + simp_or_rw [Real.pow_half_two (by arith)]; -- TODO(RFM): Technically ← but no pattern to match on otherwise. register_rewrite_map "pow_half_two-rev"; "?a" => "(pow (pow ?a 0.5) 2)" := - simp_or_rw [Real.pow_half_two (by positivity_ext)]; + simp_or_rw [Real.pow_half_two (by arith)]; register_rewrite_map "binomial_two"; "(pow (add ?a ?b) 2)" => "(add (pow ?a 2) (add (mul 2 (mul ?a ?b)) (pow ?b 2)))" := simp_or_rw [Real.binomial_two]; @@ -279,19 +279,19 @@ register_rewrite_map "exp_neg_eq_one_div-rev" ; "(div 1 (exp ?a))" => "(exp (neg simp_or_rw [←Real.exp_neg_eq_one_div]; register_rewrite_map "log_mul" ; "(log (mul ?a ?b))" => "(add (log ?a) (log ?b))" := - simp_or_rw [Real.log_mul (by positivity_ext) (by positivity_ext)]; + simp_or_rw [Real.log_mul (by arith) (by arith)]; register_rewrite_map "log_mul-rev"; "(add (log ?a) (log ?b))" => "(log (mul ?a ?b))" := - simp_or_rw [←Real.log_mul (by positivity_ext) (by positivity_ext)]; + simp_or_rw [←Real.log_mul (by arith) (by arith)]; register_rewrite_map "log_div" ; "(log (div ?a ?b))" => "(sub (log ?a) (log ?b))" := - simp_or_rw [Real.log_div (by positivity_ext) (by positivity_ext)]; + simp_or_rw [Real.log_div (by arith) (by arith)]; register_rewrite_map "log_div-rev"; "(sub (log ?a) (log ?b))" => "(log (div ?a ?b))" := - simp_or_rw [←Real.log_div (by positivity_ext) (by positivity_ext)]; + simp_or_rw [←Real.log_div (by arith) (by arith)]; register_rewrite_map "exp_log" ; "(exp (log ?a))" => "?a" := - simp_or_rw [Real.exp_log (by positivity_ext)]; + simp_or_rw [Real.exp_log (by arith)]; register_rewrite_map "log_exp" ; "(log (exp ?a))" => "?a" := simp_or_rw [Real.log_exp]; diff --git a/CvxLean/Tactic/DCP/Dcp.lean b/CvxLean/Tactic/DCP/Dcp.lean index f67785fc..eafe52cf 100644 --- a/CvxLean/Tactic/DCP/Dcp.lean +++ b/CvxLean/Tactic/DCP/Dcp.lean @@ -9,7 +9,7 @@ import CvxLean.Tactic.DCP.Tree import CvxLean.Tactic.DCP.OC import CvxLean.Meta.Util.Expr import CvxLean.Tactic.Solver.Float.OptimizationParam -import CvxLean.Tactic.Util.PositivityExt +import CvxLean.Tactic.Arith.Arith namespace CvxLean diff --git a/CvxLean/Test/Convexify/DQCP.lean b/CvxLean/Test/Convexify/DQCP.lean index 80ff7e57..76ae5e09 100644 --- a/CvxLean/Test/Convexify/DQCP.lean +++ b/CvxLean/Test/Convexify/DQCP.lean @@ -4,82 +4,6 @@ import CvxLean.Tactic.Convexify.Basic import CvxLean.Tactic.Convexify.Convexify import CvxLean.Test.Util.TimeCmd --- TODO: Move. -lemma Real.one_sub_one_div_sq_nonneg_of_le_one {x : ℝ} : - 0 < x → 0 ≤ 1 - x → 0 ≤ (1 / x ^ 2) - 1 := - fun h1 h2 => by - have hx1 : x ≤ 1 := by linarith - have h0x : 0 ≤ x := by positivity - have h0x2 : 0 < x ^ 2 := by positivity - rw [le_sub_iff_add_le, zero_add, le_div_iff h0x2, one_mul] - simpa [abs_eq_self.mpr h0x] - -lemma Real.one_sub_sq_nonneg_of_le_one {x : ℝ} : - 0 ≤ x → 0 ≤ 1 - x → 0 ≤ 1 - x ^ (2 : ℝ) := - fun h1 h2 => by - have hx1 : x ≤ 1 := by linarith - rw [le_sub_iff_add_le, zero_add] - simpa [abs_eq_self.mpr h1] - -namespace Mathlib.Meta.Positivity - -open Lean.Meta Qq - -@[positivity ((1 / (_ : ℝ) ^ (2 : ℝ)) - 1)] -def evalOneDivSqSubOne : PositivityExt where eval {_ _α} zα pα e := do - let (.app (.app _sub (.app (.app _div _one) (.app (.app _pow (x : Q(ℝ))) _two))) _one') ← - withReducible (whnf e) | throwError "not ((1 / x ^ 2) - 1)" - -- 0 < x ? - let h1 := - match ← core zα pα x with - | .positive pa => some pa - | _ => none - -- 0 ≤ 1 - x ? - let h2 := ← do - match ← core zα pα (q((1 : ℝ) - $x) : Q(ℝ)) with - | .nonnegative pa => return some pa - | .positive pa => - let pa ← mkAppM ``le_of_lt #[pa] - return some pa - | _ => return none - -- If 0 < x and 0 ≤ 1 - x, then 0 ≤ (1 / x ^ 2) - 1 - match h1, h2 with - | some h1', some h2' => - let pa' ← mkAppM ``Real.one_sub_one_div_sq_nonneg_of_le_one #[h1', h2'] - pure (.nonnegative pa') - | _, _ => - pure .none - -@[positivity (1 - ((_ : ℝ) ^ (2 : ℝ)))] -def evalOneSubSq : PositivityExt where eval {_ _α} zα pα e := do - let (.app (.app _sub _one) (.app (.app _pow (x : Q(ℝ))) _two)) ← - withReducible (whnf e) | throwError "not (1 - x ^ 2)" - -- 0 ≤ x ? - let h1 := ← do - match ← core zα pα x with - | .nonnegative pa => return some pa - | .positive pa => - let pa ← mkAppM ``le_of_lt #[pa] - return some pa - | _ => return none - -- 0 ≤ 1 - x ? - let h2 := ← do - match ← core zα pα (q((1 : ℝ) - $x) : Q(ℝ)) with - | .nonnegative pa => return some pa - | .positive pa => - let pa ← mkAppM ``le_of_lt #[pa] - return some pa - | _ => return none - -- If 0 ≤ x and 0 ≤ 1 - x, then 0 ≤ 1 - x ^ 2 - match h1, h2 with - | some h1', some h2' => - let pa' ← mkAppM ``Real.one_sub_sq_nonneg_of_le_one #[h1', h2'] - pure (.nonnegative pa') - | _, _ => - pure .none - -end Mathlib.Meta.Positivity - namespace DQCP noncomputable section diff --git a/CvxLean/Test/Convexify/Unit.lean b/CvxLean/Test/Convexify/Unit.lean index f11fe568..e27b8307 100644 --- a/CvxLean/Test/Convexify/Unit.lean +++ b/CvxLean/Test/Convexify/Unit.lean @@ -698,56 +698,6 @@ time_cmd equivalence addDivObjRed/addDivObjAuto : addDivObj := by #print addDivObjAuto -end - -end Unit - --- NOTE(RFM): This is the only way to extend the positivity tactic. - -lemma Real.exp_sub_one_pos_of_pos {x : ℝ} : 0 < x → 0 < Real.exp x - 1 := - fun h => by simp [h] - -namespace Mathlib.Meta.Positivity -open Lean.Meta Qq - -@[positivity ((Real.exp (_ : ℝ)) - 1)] -def evalExpSubOne : PositivityExt where eval {_ _α} zα pα e := do - let (.app (.app _sub (.app _exp (x : Q(ℝ)))) _one) ← - withReducible (whnf e) | throwError "not (Real.exp x - 1)" - match ← core zα pα x with - | .positive pa => - let pa' ← mkAppM ``Real.exp_sub_one_pos_of_pos #[pa] - pure (.positive pa') - | _ => - pure .none - -end Mathlib.Meta.Positivity - -lemma Real.one_sub_div_exp_pos_of_pos {x : ℝ} : 0 < x → 0 < 1 - 1 / Real.exp x := - fun h => by field_simp; positivity - -namespace Mathlib.Meta.Positivity -open Lean.Meta Qq - -@[positivity (1 - (1 / (Real.exp (_ : ℝ))))] -def evalOneSubDivExp : PositivityExt where eval {_ _α} zα pα e := do - let (.app (.app _sub _one) (.app (.app _div _one') (.app _exp (x : Q(ℝ))))) ← - withReducible (whnf e) | throwError "not (1 - 1 / Real.exp x)" - match ← core zα pα x with - | .positive pa => - let pa' ← mkAppM ``Real.one_sub_div_exp_pos_of_pos #[pa] - pure (.positive pa') - | _ => - pure .none - -end Mathlib.Meta.Positivity - -namespace Unit - -noncomputable section - -open CvxLean Minimization Real - -- add_div (constr) def addDivConstr := optimization (x y : ℝ) From 023143edb8f9e555d753923d04d18be91243503f Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 16:30:22 -0500 Subject: [PATCH 101/189] fix: `arith` imports --- CvxLean/Tactic/Arith/Arith.lean | 1 + CvxLean/Tactic/Convexify/RewriteMapLibrary.lean | 2 +- 2 files changed, 2 insertions(+), 1 deletion(-) diff --git a/CvxLean/Tactic/Arith/Arith.lean b/CvxLean/Tactic/Arith/Arith.lean index e87a7913..fed6818f 100644 --- a/CvxLean/Tactic/Arith/Arith.lean +++ b/CvxLean/Tactic/Arith/Arith.lean @@ -2,6 +2,7 @@ import Lean import Qq import Mathlib.Data.Real.Basic import Mathlib.Tactic.Positivity +import CvxLean.Tactic.Arith.PositivityExt namespace Tactic diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index ad7ca14a..aaf8ac2b 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -1,6 +1,6 @@ import CvxLean.Tactic.Convexify.RewriteMapCmd import CvxLean.Tactic.Convexify.Basic -import CvxLean.Tactic.Arith.Basic +import CvxLean.Tactic.Arith.Arith -- TODO: Move. lemma Real.log_eq_log {x y : ℝ} (hx : 0 < x) (hy : 0 < y) : Real.log x = Real.log y ↔ x = y := From 7792d257059ed4526dd6d55a7914ba912da43712 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 16:31:16 -0500 Subject: [PATCH 102/189] fix: typo in `PositivityExt.lean` --- CvxLean/Tactic/Arith/{PositvityExt.lean => PositivityExt.lean} | 0 1 file changed, 0 insertions(+), 0 deletions(-) rename CvxLean/Tactic/Arith/{PositvityExt.lean => PositivityExt.lean} (100%) diff --git a/CvxLean/Tactic/Arith/PositvityExt.lean b/CvxLean/Tactic/Arith/PositivityExt.lean similarity index 100% rename from CvxLean/Tactic/Arith/PositvityExt.lean rename to CvxLean/Tactic/Arith/PositivityExt.lean From 75879eb5bb32829e65c2ad098fbc58055321207a Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 16:32:38 -0500 Subject: [PATCH 103/189] fix: make DCP use `arith` --- CvxLean/Tactic/DCP/Dcp.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/CvxLean/Tactic/DCP/Dcp.lean b/CvxLean/Tactic/DCP/Dcp.lean index eafe52cf..d5d1c195 100644 --- a/CvxLean/Tactic/DCP/Dcp.lean +++ b/CvxLean/Tactic/DCP/Dcp.lean @@ -263,7 +263,7 @@ partial def findVConditions (originalConstrVars : Array LocalDecl) (constraints let vcondProofTy ← mkForallFVars args vcondProofTyBody let (e, _) ← Lean.Elab.Term.TermElabM.run <| Lean.Elab.Term.commitIfNoErrors? <| do - let tac ← `(by intros; try { positivity_ext <;> linarith <;> norm_num }) + let tac ← `(by intros; try { arith <;> linarith <;> norm_num }) let v ← Lean.Elab.Term.elabTerm tac.raw (some vcondProofTy) Lean.Elab.Term.synthesizeSyntheticMVarsNoPostponing instantiateMVars v From 68a46a1adfb7fe7afb34059de3806b60e816b64b Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 16:42:44 -0500 Subject: [PATCH 104/189] refactor: move real lemmas from `RewriteMapLibrary` --- CvxLean/Lib/Math/Data/Real.lean | 34 +++++++++++++ .../Tactic/Convexify/RewriteMapLibrary.lean | 51 ++++++------------- 2 files changed, 50 insertions(+), 35 deletions(-) diff --git a/CvxLean/Lib/Math/Data/Real.lean b/CvxLean/Lib/Math/Data/Real.lean index 71c87a07..6862c406 100644 --- a/CvxLean/Lib/Math/Data/Real.lean +++ b/CvxLean/Lib/Math/Data/Real.lean @@ -13,4 +13,38 @@ noncomputable def huber (x : Real) := noncomputable def klDiv (x y : Real) := x * log (x / y) - x + y +/- Lemmas used in `RewriteMapLibrary`. -/ +section Lemmas + +lemma Real.log_eq_log {x y : ℝ} (hx : 0 < x) (hy : 0 < y) : Real.log x = Real.log y ↔ x = y := + ⟨fun h => by + have hxmem := Set.mem_Ioi.2 hx + have hymem := Set.mem_Ioi.2 hy + have heq : Set.restrict (Set.Ioi 0) log ⟨x, hxmem⟩ = + Set.restrict (Set.Ioi 0) log ⟨y, hymem⟩ := by + simp [h] + have h := Real.log_injOn_pos.injective heq + simp [Subtype.eq] at h + exact h, + fun h => by rw [h]⟩ + +lemma Real.div_pow_eq_mul_pow_neg {a b c : ℝ} (hb : 0 ≤ b) : a / (b ^ c) = a * b ^ (-c) := by + rw [div_eq_mul_inv, ←rpow_neg hb] + +lemma Real.one_div_eq_pow_neg_one {a : ℝ} (ha : 0 < a) : 1 / a = a ^ (-1) := by + rw [Real.rpow_neg (le_of_lt ha), rpow_one, div_eq_mul_inv, one_mul] + +lemma Real.pow_half_two {x : ℝ} (hx : 0 ≤ x) : (x ^ (1 / 2)) ^ 2 = x := by + show Real.rpow (Real.rpow _ _) _ = _ + rw [rpow_eq_pow, rpow_eq_pow, ← rpow_mul hx] + norm_num + +lemma Real.binomial_two (x y : ℝ) : (x + y) ^ 2 = x ^ 2 + (2 * (x * y) + y ^ 2) := by + simp only [rpow_two]; ring + +lemma Real.exp_neg_eq_one_div (x : ℝ) : exp (-x) = 1 / exp x := by + rw [exp_neg, inv_eq_one_div] + +end Lemmas + end Real diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index aaf8ac2b..90d35864 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -2,41 +2,6 @@ import CvxLean.Tactic.Convexify.RewriteMapCmd import CvxLean.Tactic.Convexify.Basic import CvxLean.Tactic.Arith.Arith --- TODO: Move. -lemma Real.log_eq_log {x y : ℝ} (hx : 0 < x) (hy : 0 < y) : Real.log x = Real.log y ↔ x = y := - ⟨fun h => by - have hxmem := Set.mem_Ioi.2 hx - have hymem := Set.mem_Ioi.2 hy - have heq : Set.restrict (Set.Ioi 0) log ⟨x, hxmem⟩ = - Set.restrict (Set.Ioi 0) log ⟨y, hymem⟩ := by - simp [h] - have h := Real.log_injOn_pos.injective heq - simp [Subtype.eq] at h - exact h, - fun h => by rw [h]⟩ - --- TODO: Move. -lemma Real.div_pow_eq_mul_pow_neg {a b c : ℝ} (hb : 0 ≤ b) : a / (b ^ c) = a * b ^ (-c) := by - rw [div_eq_mul_inv, ←rpow_neg hb] - --- TODO: Move. -lemma Real.one_div_eq_pow_neg_one {a : ℝ} (ha : 0 < a) : 1 / a = a ^ (-1) := by - rw [Real.rpow_neg (le_of_lt ha), rpow_one, div_eq_mul_inv, one_mul] - --- TODO: Move. -lemma Real.pow_half_two {x : ℝ} (hx : 0 ≤ x) : (x ^ (1 / 2)) ^ 2 = x := by - show Real.rpow (Real.rpow _ _) _ = _ - rw [rpow_eq_pow, rpow_eq_pow, ← rpow_mul hx] - norm_num - --- TODO: Move. -lemma Real.binomial_two (x y : ℝ) : (x + y) ^ 2 = x ^ 2 + (2 * (x * y) + y ^ 2) := by - simp only [rpow_two]; ring - --- TODO: Move. -lemma Real.exp_neg_eq_one_div (x : ℝ) : exp (-x) = 1 / exp x := by - rw [exp_neg, inv_eq_one_div] - /-- Attempt to close the goal using the lemma specified with a combination of `simp` and `rw`. -/ syntax "simp_or_rw" "[" Lean.Parser.Tactic.rwRule "]" : tactic @@ -252,6 +217,7 @@ register_rewrite_map "pow_half_two-rev"; "?a" => "(pow (pow ?a 0.5) 2)" := register_rewrite_map "binomial_two"; "(pow (add ?a ?b) 2)" => "(add (pow ?a 2) (add (mul 2 (mul ?a ?b)) (pow ?b 2)))" := simp_or_rw [Real.binomial_two]; + /- Exponential and logarithm rules. -/ register_rewrite_map "exp_add" ; "(exp (add ?a ?b))" => "(mul (exp ?a) (exp ?b))" := @@ -297,6 +263,15 @@ register_rewrite_map "log_exp" ; "(log (exp ?a))" => "?a" := simp_or_rw [Real.log_exp]; +/- Absolute value rules. -/ + +register_rewrite_map "abs_nonneg" ; "(abs ?a)" => "?a" := + simp_or_rw [abs_eq_self.mpr (by arith)]; + +register_rewrite_map "abs_nonpos" ; "(abs ?a)" => "(neg ?a)" := + simp_or_rw [abs_eq_neg_self.mpr (by arith)]; + + /- Atom folding. -/ register_rewrite_map "xexp_fold"; "(mul ?a (exp ?a))" => "(xexp ?a)" := @@ -323,6 +298,12 @@ register_rewrite_map "geo_fold"; "(sqrt (mul ?a ?b))" => "(geo ?a ?b)" := register_rewrite_map "geo_unfold"; "(geo ?a ?b)" => "(sqrt (mul ?a ?b))" := rfl; +register_rewrite_map "lse_fold"; "(log (add (exp ?a) (exp ?b)))" => "(lse ?a ?b)" := + rfl; + +register_rewrite_map "lse_unfold"; "(lse ?a ?b)" => "(log (add (exp ?a) (exp ?b)))" := + rfl; + register_rewrite_map "norm2_fold"; "(sqrt (add (pow ?a 2) (pow ?b 2)))" => "(norm2 ?a ?b)" := rfl; From 1522d5d8d404e8da980649fbd8307cefe03b8923 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 16:49:06 -0500 Subject: [PATCH 105/189] feat: add rules for new atoms --- CvxLean/Lib/Math/Data/Real.lean | 3 +++ CvxLean/Tactic/Convexify/RewriteMapLibrary.lean | 6 ++++++ rewriter/src/rules.rs | 15 +++++++++++++++ 3 files changed, 24 insertions(+) diff --git a/CvxLean/Lib/Math/Data/Real.lean b/CvxLean/Lib/Math/Data/Real.lean index 6862c406..e544193b 100644 --- a/CvxLean/Lib/Math/Data/Real.lean +++ b/CvxLean/Lib/Math/Data/Real.lean @@ -34,6 +34,9 @@ lemma Real.div_pow_eq_mul_pow_neg {a b c : ℝ} (hb : 0 ≤ b) : a / (b ^ c) = a lemma Real.one_div_eq_pow_neg_one {a : ℝ} (ha : 0 < a) : 1 / a = a ^ (-1) := by rw [Real.rpow_neg (le_of_lt ha), rpow_one, div_eq_mul_inv, one_mul] +lemma Real.inv_eq_pow_neg_one {a : ℝ} (ha : 0 < a) : a⁻¹ = a ^ (-1) := by + rw [inv_eq_one_div, Real.one_div_eq_pow_neg_one ha] + lemma Real.pow_half_two {x : ℝ} (hx : 0 ≤ x) : (x ^ (1 / 2)) ^ 2 = x := by show Real.rpow (Real.rpow _ _) _ = _ rw [rpow_eq_pow, rpow_eq_pow, ← rpow_mul hx] diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index 90d35864..a1fe3e2a 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -147,6 +147,9 @@ register_rewrite_map "mul_div-rev" ; "(div (mul ?a ?b) ?c)" => "(mul ?a (div ?b register_rewrite_map "div_self" ; "(div ?a ?a)" => "1" := simp_or_rw [div_self (G₀ := ℝ) (by arith)]; +register_rewrite_map "inv_eq_one_div" ; "(inv ?a)" => "(div 1 ?a)" := + simp_or_rw [inv_eq_one_div (G := ℝ)]; + /- Power and square root rules. -/ @@ -217,6 +220,9 @@ register_rewrite_map "pow_half_two-rev"; "?a" => "(pow (pow ?a 0.5) 2)" := register_rewrite_map "binomial_two"; "(pow (add ?a ?b) 2)" => "(add (pow ?a 2) (add (mul 2 (mul ?a ?b)) (pow ?b 2)))" := simp_or_rw [Real.binomial_two]; +register_rewrite_map "inv_eq_pow_neg_one"; "(inv ?a)" => "(pow ?a (neg 1))" := + simp_or_rw [Real.inv_eq_pow_neg_one (by arith)]; + /- Exponential and logarithm rules. -/ diff --git a/rewriter/src/rules.rs b/rewriter/src/rules.rs index 7098c518..c454fda7 100644 --- a/rewriter/src/rules.rs +++ b/rewriter/src/rules.rs @@ -108,6 +108,8 @@ pub fn rules() -> Vec> { vec![ rw!("div_self"; "(div ?a ?a)" => "1" if is_not_zero("?a")), + rw!("inv_eq_one_div"; "(inv ?a)" => "(div 1 ?a)" if is_not_zero("?a")), + /* Power and square root rules. */ @@ -168,6 +170,8 @@ pub fn rules() -> Vec> { vec![ rw!("binomial_two"; "(pow (add ?a ?b) 2)" => "(add (pow ?a 2) (add (mul 2 (mul ?a ?b)) (pow ?b 2)))"), + rw!("inv_eq_pow_neg_one"; "(inv ?a)" => "(pow ?a (neg 1))" if is_not_zero("?a")), + /* Exponential and logarithm rules. */ @@ -204,6 +208,13 @@ pub fn rules() -> Vec> { vec![ rw!("log_exp"; "(log (exp ?a))" => "?a"), + /* Abs rules. */ + + rw!("abs_nonneg"; "(abs ?a)" => "?a" if is_ge_zero("?a")), + + rw!("abs_nonpos"; "(abs ?a)" => "(neg ?a)" if is_ge_zero("(neg ?a)")), + + /* Atom folding and unfolding rules. */ rw!("xexp_fold"; "(mul ?a (exp ?a))" => "(xexp ?a)" @@ -230,6 +241,10 @@ pub fn rules() -> Vec> { vec![ rw!("geo_unfold"; "(geo ?a ?b)" => "(sqrt (mul ?a ?b))" if is_gt_zero("?a") if is_gt_zero("?b")), + rw!("lse_fold"; "(log (add (exp ?a) (exp ?b)))" => "(lse ?a ?b)"), + + rw!("lse_unfold"; "(lse ?a ?b)" => "(log (add (exp ?a) (exp ?b)))"), + rw!("norm2_fold"; "(sqrt (add (pow ?a 2) (pow ?b 2)))" => "(norm2 ?a ?b)"), rw!("norm2_unfold"; "(norm2 ?a ?b)" => "(sqrt (add (pow ?a 2) (pow ?b 2)))"), From 5654cc2c44569f1016eaf4c058983dd8d9151331 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 16:51:14 -0500 Subject: [PATCH 106/189] feat: rule `inv_inv` --- CvxLean/Tactic/Convexify/RewriteMapLibrary.lean | 3 +++ rewriter/src/rules.rs | 2 ++ 2 files changed, 5 insertions(+) diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index a1fe3e2a..c51dd7aa 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -150,6 +150,9 @@ register_rewrite_map "div_self" ; "(div ?a ?a)" => "1" := register_rewrite_map "inv_eq_one_div" ; "(inv ?a)" => "(div 1 ?a)" := simp_or_rw [inv_eq_one_div (G := ℝ)]; +register_rewrite_map "inv_inv" ; "(inv (inv ?a))" => "?a" := + simp_or_rw [inv_inv]; + /- Power and square root rules. -/ diff --git a/rewriter/src/rules.rs b/rewriter/src/rules.rs index c454fda7..057d7a2c 100644 --- a/rewriter/src/rules.rs +++ b/rewriter/src/rules.rs @@ -110,6 +110,8 @@ pub fn rules() -> Vec> { vec![ rw!("inv_eq_one_div"; "(inv ?a)" => "(div 1 ?a)" if is_not_zero("?a")), + rw!("inv_inv"; "(inv (inv ?a))" => "?a" if is_not_zero("?a")), + /* Power and square root rules. */ From 0e69f97c9a908f6981529309ecf9cbccaf3ef2e6 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 16:57:42 -0500 Subject: [PATCH 107/189] feat: new atom `Vec.toMatrix` --- CvxLean/Lib/Math/Data/Vec.lean | 9 ++++++--- CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean | 19 ++----------------- .../DCP/AtomLibrary/Fns/VecToMatrix.lean | 17 +++++++++++++++++ 3 files changed, 25 insertions(+), 20 deletions(-) create mode 100644 CvxLean/Tactic/DCP/AtomLibrary/Fns/VecToMatrix.lean diff --git a/CvxLean/Lib/Math/Data/Vec.lean b/CvxLean/Lib/Math/Data/Vec.lean index c530fb68..acde2c40 100644 --- a/CvxLean/Lib/Math/Data/Vec.lean +++ b/CvxLean/Lib/Math/Data/Vec.lean @@ -9,6 +9,12 @@ def abs [Abs α] (x : m → α) : m → α := instance [Abs α] : Abs (m → α) := ⟨abs⟩ +def const (n : ℕ) (k : ℝ) : Fin n → ℝ := + fun _ => k + +def toMatrix {n : ℕ} (x : Fin n → ℝ) : Fin n → Fin 1 → ℝ := + fun i => ![x i] + section AddCommMonoid variable [AddCommMonoid α] {m : Nat} {n : Nat} (x : Fin m → α) (y : Fin n → α) @@ -29,9 +35,6 @@ instance : Norm (m → ℝ) where variable (x y : m → ℝ) -def const (n : ℕ) (k : ℝ) : Fin n → ℝ := - fun _ => k - def exp : m → ℝ := fun i => Real.exp (x i) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean index a042c711..09cf5094 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean @@ -1,6 +1,7 @@ import CvxLean.Tactic.DCP.Atoms import CvxLean.Tactic.DCP.AtomLibrary.Sets.Cones import CvxLean.Tactic.DCP.AtomLibrary.Fns.VecCons +import CvxLean.Tactic.DCP.AtomLibrary.Fns.VecToMatrix namespace CvxLean @@ -27,28 +28,12 @@ optimality by exact h vconditionElimination --- TODO: There is an issue matching functions so we need to wrap it. -def test (x : (Fin n) → ℝ) := fun i => ![x i] - -declare_atom vecCons2 [affine] (n : Nat)& (x : (Fin n) → ℝ)? : - test x := -bconditions -homogenity by - unfold test - ext i - simp -additivity by - unfold test - ext i - simp -optimality le_refl _ - declare_atom Vec.sq [convex] (n : ℕ)& (x : Fin n → ℝ)? : x ^ 2 := vconditions implementationVars (t : Fin n → ℝ) implementationObjective (t) implementationConstraints - (c1 : Vec.rotatedSoCone t (fun _ => 1/2) (test x)) + (c1 : Vec.rotatedSoCone t (fun _ => 1/2) (Vec.toMatrix x)) solution (t := x ^ 2) solutionEqualsAtom rfl diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/VecToMatrix.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/VecToMatrix.lean new file mode 100644 index 00000000..41b0c2ed --- /dev/null +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/VecToMatrix.lean @@ -0,0 +1,17 @@ +import CvxLean.Tactic.DCP.Atoms +import CvxLean.Lib.Math.Data.Vec + +namespace CvxLean + +open Real + +declare_atom Vec.toMatrix [affine] (n : Nat)& (x : (Fin n) → ℝ)? : + Vec.toMatrix x := +bconditions +homogenity by + unfold Vec.toMatrix; ext i; simp +additivity by + unfold Vec.toMatrix; ext i; simp +optimality le_refl _ + +end CvxLean From 3c05dab10f2a17d95707bd6a3b5a29b747ec539f Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 17:08:39 -0500 Subject: [PATCH 108/189] refactor: split basic change of variables and maps --- CvxLean/Lib/Math/Data/Real.lean | 3 ++ .../Basic.lean | 42 +-------------- CvxLean/Tactic/Convexify/Convexify.lean | 1 - CvxLean/Tactic/Convexify/MapObjFun.lean | 51 +++++++++++++++++++ .../Tactic/Convexify/RewriteMapLibrary.lean | 2 +- CvxLean/Test/Convexify/DQCP.lean | 1 - 6 files changed, 56 insertions(+), 44 deletions(-) rename CvxLean/Tactic/{Convexify => ChangeOfVariables}/Basic.lean (72%) create mode 100644 CvxLean/Tactic/Convexify/MapObjFun.lean diff --git a/CvxLean/Lib/Math/Data/Real.lean b/CvxLean/Lib/Math/Data/Real.lean index e544193b..7999395f 100644 --- a/CvxLean/Lib/Math/Data/Real.lean +++ b/CvxLean/Lib/Math/Data/Real.lean @@ -42,6 +42,9 @@ lemma Real.pow_half_two {x : ℝ} (hx : 0 ≤ x) : (x ^ (1 / 2)) ^ 2 = x := by rw [rpow_eq_pow, rpow_eq_pow, ← rpow_mul hx] norm_num +lemma Real.pow_two_le_pow_two {x y : ℝ} (hx : 0 ≤ x) (hy : 0 ≤ y) : x ^ 2 ≤ y ^ 2 ↔ x ≤ y := by + rw [rpow_two, rpow_two, sq_le_sq, abs_of_nonneg hx, abs_of_nonneg hy] + lemma Real.binomial_two (x y : ℝ) : (x + y) ^ 2 = x ^ 2 + (2 * (x * y) + y ^ 2) := by simp only [rpow_two]; ring diff --git a/CvxLean/Tactic/Convexify/Basic.lean b/CvxLean/Tactic/ChangeOfVariables/Basic.lean similarity index 72% rename from CvxLean/Tactic/Convexify/Basic.lean rename to CvxLean/Tactic/ChangeOfVariables/Basic.lean index 02d081e7..02a9c57a 100644 --- a/CvxLean/Tactic/Convexify/Basic.lean +++ b/CvxLean/Tactic/ChangeOfVariables/Basic.lean @@ -3,6 +3,7 @@ import CvxLean.Lib.Minimization import CvxLean.Lib.Equivalence import CvxLean.Tactic.Conv.ConvOpt import CvxLean.Tactic.DCP.AtomLibrary.All +import CvxLean.Lib.Math.Data.Real namespace CvxLean @@ -93,25 +94,6 @@ open Minimization Real open Lean Meta Elab Tactic Term -/-- Tactic `map_objFun_log` used to map the logarithm to the objective function -attempting to prove all the side conditions with simple tactics. -/ - -elab (name := prove_log_le_log) "prove_log_le_log" : tactic => do - let mvarId ← getMainGoal - let (_, mvarId) ← mvarId.intros - let mvarId ← mvarId.casesAnd - let mvarIds ← evalTacticAt (← - `(tactic| - try { simp only [maximizeNeg] }; - refine' (log_le_log _ _).1 _ <;> - try { assumption } <;> try { field_simp } <;> try { positivity })) mvarId - replaceMainGoal mvarIds - -macro "map_objFun_log" : tactic => - `(tactic| - apply map_objective (g := Real.log) (hg := by prove_log_le_log) <;> - dsimp only [Function.comp]) - /-- Machinery to perform the change of variables x ↦ e^u. -/ elab "prove_exp_log" : tactic => do @@ -167,28 +149,6 @@ macro "map_exp_at " i:num : tactic => dsimp only [Function.comp, ExpMapAt.exp, LogMapAt.log] <;> remove_positive_constraints) -/-- Tactic `map_objFun_sq` used to square the objective function attempting to -prove all the side conditions with simple tactics. -/ - --- TODO: Move. -lemma _root_.Real.pow_two_le_pow_two {x y : ℝ} (hx : 0 ≤ x) (hy : 0 ≤ y) : x ^ 2 ≤ y ^ 2 ↔ x ≤ y := by - rw [rpow_two, rpow_two, sq_le_sq, abs_of_nonneg hx, abs_of_nonneg hy] - -elab (name := prove_pow_two_le_pow_two) "prove_pow_two_le_pow_two" : tactic => do - let mvarId ← getMainGoal - let (_, mvarId) ← mvarId.intros - let mvarId ← mvarId.casesAnd - let mvarIds ← evalTacticAt (← - `(tactic| - rw [← Real.pow_two_le_pow_two (by positivity) (by positivity)] <;> - try { assumption } <;> try { field_simp } <;> try { positivity })) mvarId - replaceMainGoal mvarIds - -macro "map_objFun_sq" : tactic => - `(tactic| - apply map_objective (g := fun x => x ^ (2 : ℝ)) (hg := by prove_pow_two_le_pow_two) <;> - dsimp only [Function.comp]) - end Tactic end CvxLean diff --git a/CvxLean/Tactic/Convexify/Convexify.lean b/CvxLean/Tactic/Convexify/Convexify.lean index 9b76c851..d6002104 100644 --- a/CvxLean/Tactic/Convexify/Convexify.lean +++ b/CvxLean/Tactic/Convexify/Convexify.lean @@ -2,7 +2,6 @@ import Lean import Mathlib.Tactic.NormNum import CvxLean.Meta.Minimization import CvxLean.Lib.Equivalence -import CvxLean.Tactic.Convexify.Basic import CvxLean.Tactic.Convexify.RewriteMapExt import CvxLean.Tactic.Convexify.RewriteMapLibrary import CvxLean.Tactic.Convexify.Egg.All diff --git a/CvxLean/Tactic/Convexify/MapObjFun.lean b/CvxLean/Tactic/Convexify/MapObjFun.lean new file mode 100644 index 00000000..92b78f65 --- /dev/null +++ b/CvxLean/Tactic/Convexify/MapObjFun.lean @@ -0,0 +1,51 @@ +import Lean +import CvxLean.Lib.Math.Data.Real + +namespace CvxLean + +namespace Tactic + +open Real + +open Lean Meta Elab Tactic Term + +/-- Tactic `map_objFun_log` used to map the logarithm to the objective function +attempting to prove all the side conditions with simple tactics. -/ + +elab (name := prove_log_le_log) "prove_log_le_log" : tactic => do + let mvarId ← getMainGoal + let (_, mvarId) ← mvarId.intros + let mvarId ← mvarId.casesAnd + let mvarIds ← evalTacticAt (← + `(tactic| + try { simp only [maximizeNeg] }; + refine' (log_le_log _ _).1 _ <;> + try { assumption } <;> try { field_simp } <;> try { positivity })) mvarId + replaceMainGoal mvarIds + +macro "map_objFun_log" : tactic => + `(tactic| + apply map_objective (g := Real.log) (hg := by prove_log_le_log) <;> + dsimp only [Function.comp]) + +/-- Tactic `map_objFun_sq` used to square the objective function attempting to +prove all the side conditions with simple tactics. -/ + +elab (name := prove_pow_two_le_pow_two) "prove_pow_two_le_pow_two" : tactic => do + let mvarId ← getMainGoal + let (_, mvarId) ← mvarId.intros + let mvarId ← mvarId.casesAnd + let mvarIds ← evalTacticAt (← + `(tactic| + rw [←Real.pow_two_le_pow_two (by positivity) (by positivity)] <;> + try { assumption } <;> try { field_simp } <;> try { positivity })) mvarId + replaceMainGoal mvarIds + +macro "map_objFun_sq" : tactic => + `(tactic| + apply map_objective (g := fun x => x ^ (2 : ℝ)) (hg := by prove_pow_two_le_pow_two) <;> + dsimp only [Function.comp]) + +end Tactic + +end CvxLean diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index c51dd7aa..c14351bb 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -1,5 +1,5 @@ import CvxLean.Tactic.Convexify.RewriteMapCmd -import CvxLean.Tactic.Convexify.Basic +import CvxLean.Tactic.Convexify.MapObjFun import CvxLean.Tactic.Arith.Arith /-- Attempt to close the goal using the lemma specified with a combination of diff --git a/CvxLean/Test/Convexify/DQCP.lean b/CvxLean/Test/Convexify/DQCP.lean index 76ae5e09..75a8791c 100644 --- a/CvxLean/Test/Convexify/DQCP.lean +++ b/CvxLean/Test/Convexify/DQCP.lean @@ -1,6 +1,5 @@ import CvxLean.Command.Solve import CvxLean.Command.Equivalence -import CvxLean.Tactic.Convexify.Basic import CvxLean.Tactic.Convexify.Convexify import CvxLean.Test.Util.TimeCmd From 8eb7148d5a035b387231b2a10d0e14afe36efdfe Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 17:13:48 -0500 Subject: [PATCH 109/189] wip: solving hypersonic shape but wrong value --- CvxLean/Test/Convexify/DQCP.lean | 29 ++++++++++++++++++++++++++--- 1 file changed, 26 insertions(+), 3 deletions(-) diff --git a/CvxLean/Test/Convexify/DQCP.lean b/CvxLean/Test/Convexify/DQCP.lean index 75a8791c..9c9dc7e0 100644 --- a/CvxLean/Test/Convexify/DQCP.lean +++ b/CvxLean/Test/Convexify/DQCP.lean @@ -51,12 +51,35 @@ def hypersonicShapeDesign (a b : ℝ) := h2 : Δx ≤ 1 h3 : a * (1 / Δx) - (1 - b) * sqrt (1 - Δx ^ 2) ≤ 0 -time_cmd reduction redqcp2/dqcp2 : hypersonicShapeDesign 0.35 0.65 := by +time_cmd reduction redqcp2/dqcp2 : hypersonicShapeDesign 0.05 0.65 := by unfold hypersonicShapeDesign; convexify - dcp --TODO: This has 0 ≤ Δx ????? + dcp --- solve dqcp2 +solve dqcp2 + +#print dqcp2 +-- def dqcp2 : Minimization (ℝ × ℝ × ℝ × ℝ × ℝ × ℝ × ℝ × ℝ) ℝ := +-- optimization (Δx : ℝ) (t₀.0 : ℝ) (t₁.1 : ℝ) (t.2 : ℝ) (y'.3 : ℝ) (t.0.4 : ℝ) (t.5 : ℝ) (t.6 : ℝ) +-- minimize t₀.0 - 1 +-- subject to +-- _ : posOrthCone (1 - Δx) +-- _ : posOrthCone (0 - 7 / 20 * (t.2 - t.6)) +-- _ : soCone (Δx + t₁.1) ![Δx - t₁.1, 2] +-- _ : soCone (t₀.0 + 1) ![t₀.0 - 1, 2 * t₁.1] +-- _ : posOrthCone (t₀.0 - 0) +-- _ : posOrthCone (t₁.1 - 0) +-- _ : posOrthCone (Δx - 0) +-- _ : soCone (Δx + t.2) ![Δx - t.2, 2 * 1] +-- _ : posOrthCone (t.2 - 0) +-- _ : posOrthCone (Δx - t.0.4) +-- _ : expCone y'.3 1 t.0.4 +-- _ : rotatedSoCone t.5 (1 / 2) ![Δx] +-- _ : rotatedSoCone (1 - t.5) (1 / 2) ![t.6] + +-- TODO: this is wrong. +#eval dqcp2.value +#eval dqcp2.solution end QCP2 From cb77388ce4d7f62068bbf457fe359da6e316750c Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 17:14:05 -0500 Subject: [PATCH 110/189] fix: imports in `MapObjFun.lean` --- CvxLean/Tactic/Convexify/MapObjFun.lean | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/CvxLean/Tactic/Convexify/MapObjFun.lean b/CvxLean/Tactic/Convexify/MapObjFun.lean index 92b78f65..137a87f1 100644 --- a/CvxLean/Tactic/Convexify/MapObjFun.lean +++ b/CvxLean/Tactic/Convexify/MapObjFun.lean @@ -1,11 +1,14 @@ import Lean +import CvxLean.Syntax.Minimization +import CvxLean.Lib.Minimization import CvxLean.Lib.Math.Data.Real +import CvxLean.Tactic.Arith.Arith namespace CvxLean namespace Tactic -open Real +open Minimization Real open Lean Meta Elab Tactic Term From b2f8ff57224e5a764ccfd397b40c5d5b302b18c9 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 17:15:57 -0500 Subject: [PATCH 111/189] fix: imports in DGP test suites --- CvxLean/Test/Convexify/AlmostDGP.lean | 1 + CvxLean/Test/Convexify/DGP.lean | 1 + 2 files changed, 2 insertions(+) diff --git a/CvxLean/Test/Convexify/AlmostDGP.lean b/CvxLean/Test/Convexify/AlmostDGP.lean index 1a9fc408..4c49d69d 100644 --- a/CvxLean/Test/Convexify/AlmostDGP.lean +++ b/CvxLean/Test/Convexify/AlmostDGP.lean @@ -1,4 +1,5 @@ import CvxLean.Command.Solve +import CvxLean.Tactic.ChangeOfVariables.Basic import CvxLean.Tactic.Convexify.Convexify import CvxLean.Test.Util.TimeCmd diff --git a/CvxLean/Test/Convexify/DGP.lean b/CvxLean/Test/Convexify/DGP.lean index 7c6ef406..9a005577 100644 --- a/CvxLean/Test/Convexify/DGP.lean +++ b/CvxLean/Test/Convexify/DGP.lean @@ -1,4 +1,5 @@ import CvxLean.Command.Solve +import CvxLean.Tactic.ChangeOfVariables.Basic import CvxLean.Tactic.Convexify.Convexify import CvxLean.Test.Util.TimeCmd From f5597862e931ef5f9c1dcea27db56ed57cafd9ec Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 17:21:16 -0500 Subject: [PATCH 112/189] feat: print number of steps --- rewriter/src/extract.rs | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/rewriter/src/extract.rs b/rewriter/src/extract.rs index 78edd283..36d6f668 100644 --- a/rewriter/src/extract.rs +++ b/rewriter/src/extract.rs @@ -213,7 +213,10 @@ pub fn get_steps_from_string(prob_s: &str, domains_vec: Vec<(String, Domain)>, d egraph.explain_equivalence(&expr, &best); let flat_explanation : &FlatExplanation = explanation.make_flat_explanation(); - + if debug { + println!("{} steps", flat_explanation.len() - 1); + } + let mut res = Vec::new(); if best_cost.0 <= Curvature::Convex { for i in 0..flat_explanation.len() - 1 { From 2de974a5e1735bc01310b50fa19bf4a954d4701a Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 17:21:35 -0500 Subject: [PATCH 113/189] fix: check nonpositivity directly --- rewriter/src/optimization.rs | 21 +++++++++++++++++++++ rewriter/src/rules.rs | 3 ++- 2 files changed, 23 insertions(+), 1 deletion(-) diff --git a/rewriter/src/optimization.rs b/rewriter/src/optimization.rs index 48a78601..5c386831 100644 --- a/rewriter/src/optimization.rs +++ b/rewriter/src/optimization.rs @@ -236,6 +236,17 @@ pub fn is_gt_zero(var: &str) -> impl Fn(&mut EGraph, Id, &Subst) -> bool { } } +#[allow(unused)] +pub fn is_lt_zero(var: &str) -> impl Fn(&mut EGraph, Id, &Subst) -> bool { + let var = var.parse().unwrap(); + move |egraph, _, subst| { + if let Some(d) = &egraph[subst[var]].data.domain { + return domain::is_neg(d); + } + return false; + } +} + pub fn is_ge_zero(var: &str) -> impl Fn(&mut EGraph, Id, &Subst) -> bool { let var = var.parse().unwrap(); move |egraph, _, subst| { @@ -246,6 +257,16 @@ pub fn is_ge_zero(var: &str) -> impl Fn(&mut EGraph, Id, &Subst) -> bool { } } +pub fn is_le_zero(var: &str) -> impl Fn(&mut EGraph, Id, &Subst) -> bool { + let var = var.parse().unwrap(); + move |egraph, _, subst| { + if let Some(d) = &egraph[subst[var]].data.domain { + return domain::is_nonpos(d); + } + return false; + } +} + pub fn is_not_zero(var: &str) -> impl Fn(&mut EGraph, Id, &Subst) -> bool { let var = var.parse().unwrap(); move |egraph, _, subst| { diff --git a/rewriter/src/rules.rs b/rewriter/src/rules.rs index 057d7a2c..e860c68d 100644 --- a/rewriter/src/rules.rs +++ b/rewriter/src/rules.rs @@ -5,6 +5,7 @@ use optimization::Optimization as Optimization; use optimization::Meta as Meta; use optimization::is_gt_zero as is_gt_zero; use optimization::is_ge_zero as is_ge_zero; +use optimization::is_le_zero as is_le_zero; use optimization::is_not_zero as is_not_zero; pub fn rules() -> Vec> { vec![ @@ -214,7 +215,7 @@ pub fn rules() -> Vec> { vec![ rw!("abs_nonneg"; "(abs ?a)" => "?a" if is_ge_zero("?a")), - rw!("abs_nonpos"; "(abs ?a)" => "(neg ?a)" if is_ge_zero("(neg ?a)")), + rw!("abs_nonpos"; "(abs ?a)" => "(neg ?a)" if is_le_zero("?a")), /* Atom folding and unfolding rules. */ From 645bab0d639e9e30d1b61c4e7496cc9541828bac Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 17:39:54 -0500 Subject: [PATCH 114/189] feat: `is_nat` --- rewriter/src/optimization.rs | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/rewriter/src/optimization.rs b/rewriter/src/optimization.rs index 5c386831..dd60bbe6 100644 --- a/rewriter/src/optimization.rs +++ b/rewriter/src/optimization.rs @@ -276,3 +276,15 @@ pub fn is_not_zero(var: &str) -> impl Fn(&mut EGraph, Id, &Subst) -> bool { return false; } } + +pub fn is_nat(var: &str) -> impl Fn(&mut EGraph, Id, &Subst) -> bool { + let var = var.parse().unwrap(); + move |egraph, _, subst| { + if let Some(d) = &egraph[subst[var]].data.domain { + if let Some(c) = Domain::get_constant(d) { + return ((c as u32) as f64) == c; + } + } + return false; + } +} From bcfed5f63161e4afd3b200b42cd574523fa02d1c Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 17:40:07 -0500 Subject: [PATCH 115/189] feat: logarithm power rules --- CvxLean/Tactic/Convexify/RewriteMapLibrary.lean | 13 +++++++++++++ rewriter/src/rules.rs | 13 +++++++++++++ 2 files changed, 26 insertions(+) diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index c14351bb..35eccdb7 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -265,6 +265,19 @@ register_rewrite_map "log_div" ; "(log (div ?a ?b))" => "(sub (log ?a) (log ?b)) register_rewrite_map "log_div-rev"; "(sub (log ?a) (log ?b))" => "(log (div ?a ?b))" := simp_or_rw [←Real.log_div (by arith) (by arith)]; +register_rewrite_map "log_pow" ; "(log (pow ?a ?b))" => "(mul ?b (log ?a))" := + simp_or_rw [Real.log_rpow (by arith)]; + +register_rewrite_map "log_pow-rev" ; "(mul ?b (log ?a))" => "(log (pow ?a ?b))" := + simp_or_rw [←Real.log_rpow (by arith)]; + +-- NOTE: Special rule that only works if the exponent is a natural number. +register_rewrite_map "log_pow_nat" ; "(log (pow ?a ?b))" => "(mul ?b (log ?a))" := + (norm_cast; simp_or_rw [Real.log_pow]); + +register_rewrite_map "log_pow_nat-rev" ; "(mul ?b (log ?a))" => "(log (pow ?a ?b))" := + (norm_cast; simp_or_rw [←Real.log_pow]); + register_rewrite_map "exp_log" ; "(exp (log ?a))" => "?a" := simp_or_rw [Real.exp_log (by arith)]; diff --git a/rewriter/src/rules.rs b/rewriter/src/rules.rs index e860c68d..f4655ec1 100644 --- a/rewriter/src/rules.rs +++ b/rewriter/src/rules.rs @@ -7,6 +7,7 @@ use optimization::is_gt_zero as is_gt_zero; use optimization::is_ge_zero as is_ge_zero; use optimization::is_le_zero as is_le_zero; use optimization::is_not_zero as is_not_zero; +use optimization::is_nat as is_nat; pub fn rules() -> Vec> { vec![ @@ -205,6 +206,18 @@ pub fn rules() -> Vec> { vec![ rw!("log_div-rev"; "(sub (log ?a) (log ?b))" => "(log (div ?a ?b))" if is_gt_zero("?a") if is_gt_zero("?b")), + + rw!("log_pow"; "(log (pow ?a ?b))" => "(mul ?b (log ?a))" + if is_gt_zero("?a")), + + rw!("log_pow-rev"; "(mul ?b (log ?a))" => "(log (pow ?a ?b))" + if is_gt_zero("?a")), + + rw!("log_pow_nat"; "(log (pow ?a ?b))" => "(mul ?b (log ?a))" + if is_nat("?b")), + + rw!("log_pow_nat-rev"; "(mul ?b (log ?a))" => "(log (pow ?a ?b))" + if is_nat("?b")), rw!("exp_log"; "(exp (log ?a))" => "?a" if is_gt_zero("?a")), From d9b9cceb64817862ba5f68d3b4716ffe0bbdfdf7 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 17:46:45 -0500 Subject: [PATCH 116/189] test: examples from DCP quiz --- rewriter/tests/test_quiz.rs | 66 +++++++++++++++++++++++++++++++++++-- 1 file changed, 64 insertions(+), 2 deletions(-) diff --git a/rewriter/tests/test_quiz.rs b/rewriter/tests/test_quiz.rs index 0c4a1039..0fbbf6f3 100644 --- a/rewriter/tests/test_quiz.rs +++ b/rewriter/tests/test_quiz.rs @@ -1,5 +1,5 @@ /*! -Tests from random (non-DCP) problems generated by the DCP quiz: +Tests adapted from random (non-DCP) problems generated by the DCP quiz: https://dcp.stanford.edu/quiz !*/ @@ -10,6 +10,68 @@ use egg_convexify::test_util::{*}; #[test] fn test_quiz_1() { convexify_check_expression_with_domain( - vec![("x", domain::free_dom())], + vec![("x", domain::pos_dom())], "(inv (inv (var x)))"); } + +#[test] +fn test_quiz_2() { + convexify_check_expression_with_domain( + vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], + "(neg (lse (log (var x)) (log (var y))))"); +} + +#[test] +fn test_quiz_3() { + convexify_check_expression_with_domain( + vec![("x", domain::nonneg_dom())], + "(pow (sqrt (var x)) 2)"); +} + +#[test] +fn test_quiz_4() { + convexify_check_expression_with_domain( + vec![("x", domain::nonneg_dom())], + "(neg (abs (sqrt (abs (var x)))))"); +} + +#[test] +fn test_quiz_5() { + convexify_check_expression_with_domain( + vec![("x", domain::free_dom())], + "(div 1 (exp (var x)))"); +} + +#[test] +fn test_quiz_6() { + convexify_check_expression_with_domain( + vec![("x", domain::nonneg_dom())], + "(neg (log (pow (mul 364 (var x)) 2)))"); +} + +#[test] +fn test_quiz_7() { + convexify_check_expression_with_domain( + vec![("x", domain::pos_dom())], + "(pow (geo (add (var x) 2) (div 1 (var x))) 2)"); +} + +#[test] +fn test_quiz_8() { + convexify_check_expression_with_domain( + vec![("x", domain::nonneg_dom())], + "(neg (log (abs (var x))))"); +} + +#[test] +fn test_quiz_9() { + convexify_check_expression_with_domain( + vec![("x", domain::pos_dom()), ("y", domain::pos_dom())], + "(div 1 (qol (inv (var x)) (inv (var y)))))"); +} + +#[test] +fn test_quiz_10() { + convexify_check_expression( + "(pow (log (exp (var x))) 2)"); +} From cd444e6777a3f0fcdac1184d826665bc012fa527 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 17:49:16 -0500 Subject: [PATCH 117/189] refactor: `rewriter` -> `egg-convexify` --- .gitignore | 4 ++-- CvxLean/Tactic/Convexify/Egg/Runner.lean | 2 +- {rewriter => egg-convexify}/.gitignore | 0 {rewriter => egg-convexify}/Cargo.lock | 0 {rewriter => egg-convexify}/Cargo.toml | 0 {rewriter => egg-convexify}/src/cost.rs | 0 {rewriter => egg-convexify}/src/curvature.rs | 0 {rewriter => egg-convexify}/src/domain.rs | 0 {rewriter => egg-convexify}/src/explain_util.rs | 0 {rewriter => egg-convexify}/src/extract.rs | 0 {rewriter => egg-convexify}/src/lib.rs | 0 {rewriter => egg-convexify}/src/main.rs | 0 {rewriter => egg-convexify}/src/optimization.rs | 0 {rewriter => egg-convexify}/src/rules.rs | 0 {rewriter => egg-convexify}/src/test_util.rs | 0 {rewriter => egg-convexify}/tests/test_almost_dgp.rs | 0 {rewriter => egg-convexify}/tests/test_cost_function.rs | 0 {rewriter => egg-convexify}/tests/test_domain.rs | 0 {rewriter => egg-convexify}/tests/test_dqcp.rs | 0 {rewriter => egg-convexify}/tests/test_gp.rs | 0 {rewriter => egg-convexify}/tests/test_main_example.rs | 0 {rewriter => egg-convexify}/tests/test_misc.rs | 0 {rewriter => egg-convexify}/tests/test_quiz.rs | 0 {rewriter => egg-convexify}/tests/test_rules.rs | 0 {rewriter => egg-convexify}/tests/test_stanford.rs | 0 lakefile.lean | 6 +++--- 26 files changed, 6 insertions(+), 6 deletions(-) rename {rewriter => egg-convexify}/.gitignore (100%) rename {rewriter => egg-convexify}/Cargo.lock (100%) rename {rewriter => egg-convexify}/Cargo.toml (100%) rename {rewriter => egg-convexify}/src/cost.rs (100%) rename {rewriter => egg-convexify}/src/curvature.rs (100%) rename {rewriter => egg-convexify}/src/domain.rs (100%) rename {rewriter => egg-convexify}/src/explain_util.rs (100%) rename {rewriter => egg-convexify}/src/extract.rs (100%) rename {rewriter => egg-convexify}/src/lib.rs (100%) rename {rewriter => egg-convexify}/src/main.rs (100%) rename {rewriter => egg-convexify}/src/optimization.rs (100%) rename {rewriter => egg-convexify}/src/rules.rs (100%) rename {rewriter => egg-convexify}/src/test_util.rs (100%) rename {rewriter => egg-convexify}/tests/test_almost_dgp.rs (100%) rename {rewriter => egg-convexify}/tests/test_cost_function.rs (100%) rename {rewriter => egg-convexify}/tests/test_domain.rs (100%) rename {rewriter => egg-convexify}/tests/test_dqcp.rs (100%) rename {rewriter => egg-convexify}/tests/test_gp.rs (100%) rename {rewriter => egg-convexify}/tests/test_main_example.rs (100%) rename {rewriter => egg-convexify}/tests/test_misc.rs (100%) rename {rewriter => egg-convexify}/tests/test_quiz.rs (100%) rename {rewriter => egg-convexify}/tests/test_rules.rs (100%) rename {rewriter => egg-convexify}/tests/test_stanford.rs (100%) diff --git a/.gitignore b/.gitignore index 979e6119..08378766 100644 --- a/.gitignore +++ b/.gitignore @@ -3,8 +3,8 @@ /build /solver/*.cbf /solver/*.sol -/rewriter/*.dot -/rewriter/*.svg +/egg-convexify/*.dot +/egg-convexify/*.svg /ffi/build /ffi/lake-packages /lake-packages diff --git a/CvxLean/Tactic/Convexify/Egg/Runner.lean b/CvxLean/Tactic/Convexify/Egg/Runner.lean index 8e1fe199..c2e4910a 100644 --- a/CvxLean/Tactic/Convexify/Egg/Runner.lean +++ b/CvxLean/Tactic/Convexify/Egg/Runner.lean @@ -101,7 +101,7 @@ namespace CvxLean def runEggRequestRaw (requestJson : String) : MetaM String := do let eggProcess ← IO.Process.spawn { - cmd := "rewriter/utils/egg-convexify", + cmd := "egg-convexify/utils/egg-convexify", stdout := IO.Process.Stdio.piped, stdin := IO.Process.Stdio.piped, stderr := IO.Process.Stdio.null diff --git a/rewriter/.gitignore b/egg-convexify/.gitignore similarity index 100% rename from rewriter/.gitignore rename to egg-convexify/.gitignore diff --git a/rewriter/Cargo.lock b/egg-convexify/Cargo.lock similarity index 100% rename from rewriter/Cargo.lock rename to egg-convexify/Cargo.lock diff --git a/rewriter/Cargo.toml b/egg-convexify/Cargo.toml similarity index 100% rename from rewriter/Cargo.toml rename to egg-convexify/Cargo.toml diff --git a/rewriter/src/cost.rs b/egg-convexify/src/cost.rs similarity index 100% rename from rewriter/src/cost.rs rename to egg-convexify/src/cost.rs diff --git a/rewriter/src/curvature.rs b/egg-convexify/src/curvature.rs similarity index 100% rename from rewriter/src/curvature.rs rename to egg-convexify/src/curvature.rs diff --git a/rewriter/src/domain.rs b/egg-convexify/src/domain.rs similarity index 100% rename from rewriter/src/domain.rs rename to egg-convexify/src/domain.rs diff --git a/rewriter/src/explain_util.rs b/egg-convexify/src/explain_util.rs similarity index 100% rename from rewriter/src/explain_util.rs rename to egg-convexify/src/explain_util.rs diff --git a/rewriter/src/extract.rs b/egg-convexify/src/extract.rs similarity index 100% rename from rewriter/src/extract.rs rename to egg-convexify/src/extract.rs diff --git a/rewriter/src/lib.rs b/egg-convexify/src/lib.rs similarity index 100% rename from rewriter/src/lib.rs rename to egg-convexify/src/lib.rs diff --git a/rewriter/src/main.rs b/egg-convexify/src/main.rs similarity index 100% rename from rewriter/src/main.rs rename to egg-convexify/src/main.rs diff --git a/rewriter/src/optimization.rs b/egg-convexify/src/optimization.rs similarity index 100% rename from rewriter/src/optimization.rs rename to egg-convexify/src/optimization.rs diff --git a/rewriter/src/rules.rs b/egg-convexify/src/rules.rs similarity index 100% rename from rewriter/src/rules.rs rename to egg-convexify/src/rules.rs diff --git a/rewriter/src/test_util.rs b/egg-convexify/src/test_util.rs similarity index 100% rename from rewriter/src/test_util.rs rename to egg-convexify/src/test_util.rs diff --git a/rewriter/tests/test_almost_dgp.rs b/egg-convexify/tests/test_almost_dgp.rs similarity index 100% rename from rewriter/tests/test_almost_dgp.rs rename to egg-convexify/tests/test_almost_dgp.rs diff --git a/rewriter/tests/test_cost_function.rs b/egg-convexify/tests/test_cost_function.rs similarity index 100% rename from rewriter/tests/test_cost_function.rs rename to egg-convexify/tests/test_cost_function.rs diff --git a/rewriter/tests/test_domain.rs b/egg-convexify/tests/test_domain.rs similarity index 100% rename from rewriter/tests/test_domain.rs rename to egg-convexify/tests/test_domain.rs diff --git a/rewriter/tests/test_dqcp.rs b/egg-convexify/tests/test_dqcp.rs similarity index 100% rename from rewriter/tests/test_dqcp.rs rename to egg-convexify/tests/test_dqcp.rs diff --git a/rewriter/tests/test_gp.rs b/egg-convexify/tests/test_gp.rs similarity index 100% rename from rewriter/tests/test_gp.rs rename to egg-convexify/tests/test_gp.rs diff --git a/rewriter/tests/test_main_example.rs b/egg-convexify/tests/test_main_example.rs similarity index 100% rename from rewriter/tests/test_main_example.rs rename to egg-convexify/tests/test_main_example.rs diff --git a/rewriter/tests/test_misc.rs b/egg-convexify/tests/test_misc.rs similarity index 100% rename from rewriter/tests/test_misc.rs rename to egg-convexify/tests/test_misc.rs diff --git a/rewriter/tests/test_quiz.rs b/egg-convexify/tests/test_quiz.rs similarity index 100% rename from rewriter/tests/test_quiz.rs rename to egg-convexify/tests/test_quiz.rs diff --git a/rewriter/tests/test_rules.rs b/egg-convexify/tests/test_rules.rs similarity index 100% rename from rewriter/tests/test_rules.rs rename to egg-convexify/tests/test_rules.rs diff --git a/rewriter/tests/test_stanford.rs b/egg-convexify/tests/test_stanford.rs similarity index 100% rename from rewriter/tests/test_stanford.rs rename to egg-convexify/tests/test_stanford.rs diff --git a/lakefile.lean b/lakefile.lean index b3c9999a..490d476e 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -44,15 +44,15 @@ SchedulerM (BuildJob FilePath) := @[default_target] target EggConvexify (pkg) : FilePath := do - let buildDir := pkg.dir / "rewriter" + let buildDir := pkg.dir / "egg-convexify" let binFile := buildDir / "target" / "release" / "egg-convexify" let dest := buildDir / "utils" / "egg-convexify" let manifestFile := buildDir / "Cargo.toml" buildCargo binFile manifestFile dest #[] script EggClean := do - let targetDir : FilePath := "." / "rewriter" / "target" - let utilsDir : FilePath := "." / "rewriter" / "utils" + let targetDir : FilePath := "." / "egg-convexify" / "target" + let utilsDir : FilePath := "." / "egg-convexify" / "utils" let out ← IO.Process.output { cmd := "rm" From 6de49ad5e78621a8d2170e5ef3a591aace9566a2 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 18:03:50 -0500 Subject: [PATCH 118/189] refactor: move real instances --- CvxLean/Lib/Math/Data/Real.lean | 29 +++++++++++++++++++++++++---- 1 file changed, 25 insertions(+), 4 deletions(-) diff --git a/CvxLean/Lib/Math/Data/Real.lean b/CvxLean/Lib/Math/Data/Real.lean index 7999395f..d463a3c8 100644 --- a/CvxLean/Lib/Math/Data/Real.lean +++ b/CvxLean/Lib/Math/Data/Real.lean @@ -4,6 +4,9 @@ import Mathlib.Analysis.SpecialFunctions.Pow.Real namespace Real +/- New named functions. Each corresponds to an atom. -/ +section Functions + noncomputable def entr (x : Real) := -(x * Real.log x) @@ -13,10 +16,25 @@ noncomputable def huber (x : Real) := noncomputable def klDiv (x y : Real) := x * log (x / y) - x + y +end Functions + +/- Useful to construct expressions, as otherwise they are derived from other +instances. -/ +section Instances + +noncomputable instance Real.instDivReal : Div ℝ := + by infer_instance + +noncomputable instance Real.instAbsReal : Abs ℝ := + by infer_instance + +end Instances + /- Lemmas used in `RewriteMapLibrary`. -/ section Lemmas -lemma Real.log_eq_log {x y : ℝ} (hx : 0 < x) (hy : 0 < y) : Real.log x = Real.log y ↔ x = y := +lemma Real.log_eq_log {x y : ℝ} (hx : 0 < x) (hy : 0 < y) : + Real.log x = Real.log y ↔ x = y := ⟨fun h => by have hxmem := Set.mem_Ioi.2 hx have hymem := Set.mem_Ioi.2 hy @@ -28,7 +46,8 @@ lemma Real.log_eq_log {x y : ℝ} (hx : 0 < x) (hy : 0 < y) : Real.log x = Real. exact h, fun h => by rw [h]⟩ -lemma Real.div_pow_eq_mul_pow_neg {a b c : ℝ} (hb : 0 ≤ b) : a / (b ^ c) = a * b ^ (-c) := by +lemma Real.div_pow_eq_mul_pow_neg {a b c : ℝ} (hb : 0 ≤ b) : + a / (b ^ c) = a * b ^ (-c) := by rw [div_eq_mul_inv, ←rpow_neg hb] lemma Real.one_div_eq_pow_neg_one {a : ℝ} (ha : 0 < a) : 1 / a = a ^ (-1) := by @@ -42,10 +61,12 @@ lemma Real.pow_half_two {x : ℝ} (hx : 0 ≤ x) : (x ^ (1 / 2)) ^ 2 = x := by rw [rpow_eq_pow, rpow_eq_pow, ← rpow_mul hx] norm_num -lemma Real.pow_two_le_pow_two {x y : ℝ} (hx : 0 ≤ x) (hy : 0 ≤ y) : x ^ 2 ≤ y ^ 2 ↔ x ≤ y := by +lemma Real.pow_two_le_pow_two {x y : ℝ} (hx : 0 ≤ x) (hy : 0 ≤ y) : + x ^ 2 ≤ y ^ 2 ↔ x ≤ y := by rw [rpow_two, rpow_two, sq_le_sq, abs_of_nonneg hx, abs_of_nonneg hy] -lemma Real.binomial_two (x y : ℝ) : (x + y) ^ 2 = x ^ 2 + (2 * (x * y) + y ^ 2) := by +lemma Real.binomial_two (x y : ℝ) : + (x + y) ^ 2 = x ^ 2 + (2 * (x * y) + y ^ 2) := by simp only [rpow_two]; ring lemma Real.exp_neg_eq_one_div (x : ℝ) : exp (-x) = 1 / exp x := by From 88b4a0a70b2da973d78fab8b51651bd08960a56f Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 18:07:29 -0500 Subject: [PATCH 119/189] feat: Lean-egg interface for new atoms --- CvxLean/Lib/Math/Data/Real.lean | 14 +++++--- .../Convexify/Egg/FromMinimization.lean | 28 +++++++++------ CvxLean/Tactic/Convexify/Egg/ToExpr.lean | 34 +++++++++++++++++-- 3 files changed, 58 insertions(+), 18 deletions(-) diff --git a/CvxLean/Lib/Math/Data/Real.lean b/CvxLean/Lib/Math/Data/Real.lean index d463a3c8..766906f0 100644 --- a/CvxLean/Lib/Math/Data/Real.lean +++ b/CvxLean/Lib/Math/Data/Real.lean @@ -22,11 +22,17 @@ end Functions instances. -/ section Instances -noncomputable instance Real.instDivReal : Div ℝ := - by infer_instance +noncomputable instance instDivReal : Div ℝ := by + infer_instance -noncomputable instance Real.instAbsReal : Abs ℝ := - by infer_instance +noncomputable instance instAbsReal : Abs ℝ := by + infer_instance + +noncomputable instance instMinReal : Min ℝ := by + infer_instance + +noncomputable instance instMaxReal : Max ℝ := by + infer_instance end Instances diff --git a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean index ce661141..fa1fe9ff 100644 --- a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean +++ b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean @@ -48,20 +48,26 @@ def _root_.EggTree.opMap : HashMap String (String × Nat × Array String) := ("le", ("le", 2, #[])), ("lt", ("lt", 2, #[])), ("neg", ("neg", 1, #[])), + ("inv", ("inv", 1, #[])), + ("abs", ("abs", 1, #[])), ("sqrt'", ("sqrt", 1, #[])), + ("sq", ("pow", 1, #["2"])), + ("powNegOne", ("pow", 1, #["-1"])), + ("powNegTwo", ("pow", 1, #["-2"])), ("log", ("log", 1, #[])), ("exp", ("exp", 1, #[])), ("xexp", ("xexp", 1, #[])), ("entr", ("entr", 1, #[])), + ("min", ("min", 2, #[])), + ("max", ("max", 2, #[])), ("add", ("add", 2, #[])), ("sub", ("sub", 2, #[])), ("mul1", ("mul", 2, #[])), ("mul2", ("mul", 2, #[])), - ("sq", ("pow", 1, #["2"])), - ("sqrt", ("sqrt", 1, #[])), ("div", ("div", 2, #[])), ("quadOverLin", ("qol", 2, #[])), - ("geoMean", ("qol", 2, #[])), + ("geoMean", ("geo", 2, #[])), + ("logSumExp", ("lse", 2, #[])), ("norm2₂", ("norm2", 2, #[])) ] @@ -136,24 +142,24 @@ def _root_.ExtendedEggTree.fromMinimization (e : Meta.MinimizationExpr) (vars : | Tree.node "le" #[(tl : EggTree), (tr : EggTree)] => match (tl.getVariableName? vars, tr.getNumericalValue?) with | (some v, some f) => - -- v ∈ [-inf, f] - some (h, v, ⟨"-inf", f.toString, "0", "0"⟩) + -- v ∈ (-inf, f] + some (h, v, ⟨"-inf", f.toString, "1", "0"⟩) | _ => match (tl.getNumericalValue?, tr.getVariableName? vars) with | (some f, some v) => - -- v ∈ [f, inf] - some (h, v, ⟨f.toString, "inf", "0", "0"⟩) + -- v ∈ [f, inf) + some (h, v, ⟨f.toString, "inf", "0", "1"⟩) | _ => none | Tree.node "lt" #[(tl : EggTree), (tr : EggTree)] => match (tl.getVariableName? vars, tr.getNumericalValue?) with | (some v, some f) => - -- v ∈ [-inf, f) - some (h, v, ⟨"-inf", f.toString, "0", "1"⟩) + -- v ∈ (-inf, f) + some (h, v, ⟨"-inf", f.toString, "1", "1"⟩) | _ => match (tl.getNumericalValue?, tr.getVariableName? vars) with | (some f, some v) => - -- v ∈ (f, inf] - some (h, v, ⟨f.toString, "inf", "1", "0"⟩) + -- v ∈ (f, inf) + some (h, v, ⟨f.toString, "inf", "1", "1"⟩) | _ => none | Tree.node "eq" #[(tl : EggTree), (tr : EggTree)] => match (tl.getVariableName? vars, tr.getNumericalValue?) with diff --git a/CvxLean/Tactic/Convexify/Egg/ToExpr.lean b/CvxLean/Tactic/Convexify/Egg/ToExpr.lean index 4024cc64..acc2bfab 100644 --- a/CvxLean/Tactic/Convexify/Egg/ToExpr.lean +++ b/CvxLean/Tactic/Convexify/Egg/ToExpr.lean @@ -27,9 +27,6 @@ def EggString.toEggTree (s : String) : MetaM (Tree String String) := do | Except.ok sexpr => Sexpr.toEggTree sexpr | Except.error e => throwError s!"{e}" -noncomputable instance Real.instDivReal : Div ℝ := - by infer_instance - /-- -/ partial def EggTree.toExpr (vars : List String) : Tree String String → MetaM Expr -- Numbers. @@ -81,6 +78,18 @@ partial def EggTree.toExpr (vars : List String) : Tree String String → MetaM E return mkAppN (mkConst ``Neg.neg [levelZero]) #[(mkConst ``Real), (mkConst ``Real.instNegReal), t] + -- Inverse. + | Tree.node "inv" #[t] => do + let t ← toExpr vars t + return mkAppN + (mkConst ``Inv.inv [levelZero]) + #[(mkConst ``Real), (mkConst ``Real.instInvReal), t] + -- Absolute value. + | Tree.node "abs" #[t] => do + let t ← toExpr vars t + return mkAppN + (mkConst ``Abs.abs [levelZero]) + #[(mkConst ``Real), (mkConst ``Real.instAbsReal), t] -- Square root. | Tree.node "sqrt" #[t] => do let t ← toExpr vars t @@ -99,6 +108,20 @@ partial def EggTree.toExpr (vars : List String) : Tree String String → MetaM E -- Entr. | Tree.node "entr" #[t] => EggTree.toExpr vars (Tree.node "neg" #[Tree.node "mul" #[t, Tree.node "log" #[t]]]) + -- Min. + | Tree.node "min" #[t1, t2] => do + let t1 ← toExpr vars t1 + let t2 ← toExpr vars t2 + return mkAppN + (mkConst ``Min.min [levelZero]) + #[(mkConst ``Real), (mkConst ``Real.instMinReal), t1, t2] + -- Max. + | Tree.node "max" #[t1, t2] => do + let t1 ← toExpr vars t1 + let t2 ← toExpr vars t2 + return mkAppN + (mkConst ``Max.max [levelZero]) + #[(mkConst ``Real), (mkConst ``Real.instMaxReal), t1, t2] -- Addition. | Tree.node "add" #[t1, t2] => do let t1 ← toExpr vars t1 @@ -130,6 +153,11 @@ partial def EggTree.toExpr (vars : List String) : Tree String String → MetaM E -- Geo mean. | Tree.node "geo" #[t1, t2] => EggTree.toExpr vars (Tree.node "sqrt" #[Tree.node "mul" #[t1, t2]]) + -- Log sum exp. + | Tree.node "lse" #[t1, t2] => + EggTree.toExpr vars (Tree.node "log" #[Tree.node "add" #[ + Tree.node "exp" #[t1], + Tree.node "exp" #[t2]]]) -- Norm2. | Tree.node "norm2" #[t1, t2] => EggTree.toExpr vars (Tree.node "sqrt" #[Tree.node "add" #[ From 8a1cdc8629cd712a42373f988f66fdf55a63be99 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 18:17:40 -0500 Subject: [PATCH 120/189] wip: add sub-expressions in egg steps for targeted rewrites --- CvxLean/Tactic/Convexify/Egg/Runner.lean | 8 ++++++++ egg-convexify/src/extract.rs | 15 ++++++++++----- 2 files changed, 18 insertions(+), 5 deletions(-) diff --git a/CvxLean/Tactic/Convexify/Egg/Runner.lean b/CvxLean/Tactic/Convexify/Egg/Runner.lean index c2e4910a..82b5f8a5 100644 --- a/CvxLean/Tactic/Convexify/Egg/Runner.lean +++ b/CvxLean/Tactic/Convexify/Egg/Runner.lean @@ -84,6 +84,8 @@ structure EggRewrite where rewriteName : String direction : EggRewriteDirection location : String + subexprFrom : String + subexprTo : String expectedTerm : String def EggRewrite.toString (e : EggRewrite) : String := @@ -91,6 +93,8 @@ def EggRewrite.toString (e : EggRewrite) : String := surroundQuotes "rewrite_name" ++ " : " ++ surroundQuotes e.rewriteName ++ ", " ++ surroundQuotes "direction" ++ " : " ++ surroundQuotes (e.direction.toString) ++ ", " ++ surroundQuotes "location" ++ " : " ++ surroundQuotes e.location ++ ", " ++ + surroundQuotes "subexpr_from" ++ " : " ++ surroundQuotes e.subexprFrom ++ ", " ++ + surroundQuotes "subexpr_to" ++ " : " ++ surroundQuotes e.subexprTo ++ ", " ++ surroundQuotes "expected_term" ++ " : " ++ surroundQuotes e.expectedTerm ++ "}" @@ -142,10 +146,14 @@ def parseEggResponse (responseString : String) : MetaM (Array EggRewrite) := do | "Backward" => EggRewriteDirection.Backward | _ => panic! "Unexpected rewrite direction." let location := (step.getObjValD "location").getStr! + let subexprFrom := (step.getObjValD "subexpr_from").getStr! + let subexprTo := (step.getObjValD "subexpr_to").getStr! let expectedTerm := (step.getObjValD "expected_term").getStr! { rewriteName := rewriteName, direction := direction, location := location, + subexprFrom := subexprFrom, + subexprTo := subexprTo, expectedTerm := expectedTerm } return res diff --git a/egg-convexify/src/extract.rs b/egg-convexify/src/extract.rs index 36d6f668..fffd5f2b 100644 --- a/egg-convexify/src/extract.rs +++ b/egg-convexify/src/extract.rs @@ -30,6 +30,8 @@ pub struct Step { rewrite_name : String, direction : Direction, location : String, + subexpr_from : String, + subexpr_to : String, expected_term : String, } @@ -66,11 +68,15 @@ fn get_step_aux( } if let Some(rule_name) = &next.backward_rule { + let subexpr_from = current.get_recexpr().to_string(); + let subexpr_to = next.get_recexpr().to_string(); if expected_term.is_some() { return Some(Step { rewrite_name: rule_name.to_string(), direction: Direction::Backward, location: location.clone().unwrap_or_default(), + subexpr_from: subexpr_from, + subexpr_to: subexpr_to, expected_term: expected_term.clone().unwrap(), }); } else { @@ -79,16 +85,15 @@ fn get_step_aux( } if let Some(rule_name) = &next.forward_rule { - // NOTE: Could be useful for more targetted rewrites. - // let curr_s_t: String = current.get_recexpr().to_string(); - // let next_s_t = next.get_recexpr().to_string(); - // println!("curr {} : {}", rule_name, curr_s_t); - // println!("next {} : {}", rule_name, next_s_t); + let subexpr_from = current.get_recexpr().to_string(); + let subexpr_to = next.get_recexpr().to_string(); if expected_term.is_some() { return Some(Step { rewrite_name: rule_name.to_string(), direction: Direction::Forward, location: location.clone().unwrap_or_default(), + subexpr_from: subexpr_from, + subexpr_to: subexpr_to, expected_term: expected_term.clone().unwrap(), }); } else { From df2304890549dc54bdd41e4022562ce09c781a54 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Thu, 7 Dec 2023 18:23:23 -0500 Subject: [PATCH 121/189] fix: correct name for `sqrt` atom --- CvxLean/Tactic/Convexify/Egg/FromMinimization.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean index fa1fe9ff..fdd6a74d 100644 --- a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean +++ b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean @@ -50,7 +50,7 @@ def _root_.EggTree.opMap : HashMap String (String × Nat × Array String) := ("neg", ("neg", 1, #[])), ("inv", ("inv", 1, #[])), ("abs", ("abs", 1, #[])), - ("sqrt'", ("sqrt", 1, #[])), + ("sqrt", ("sqrt", 1, #[])), ("sq", ("pow", 1, #["2"])), ("powNegOne", ("pow", 1, #["-1"])), ("powNegTwo", ("pow", 1, #["-2"])), From a8427df0585a85ef3513d39d0e2dc3cf2df107f6 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 10:08:55 -0500 Subject: [PATCH 122/189] feat: recover all convexify unit tests --- CvxLean/Lib/Cones/SOCone.lean | 3 ++- CvxLean/Tactic/Convexify/Convexify.lean | 12 +++++------- CvxLean/Tactic/Convexify/Egg/FromMinimization.lean | 1 + CvxLean/Tactic/Convexify/RewriteMapLibrary.lean | 4 ++-- CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean | 10 ++++++++++ CvxLean/Test/Convexify/Unit.lean | 6 +++--- 6 files changed, 23 insertions(+), 13 deletions(-) diff --git a/CvxLean/Lib/Cones/SOCone.lean b/CvxLean/Lib/Cones/SOCone.lean index cb44e0b4..ef2898d5 100644 --- a/CvxLean/Lib/Cones/SOCone.lean +++ b/CvxLean/Lib/Cones/SOCone.lean @@ -43,7 +43,8 @@ section Lemmas /-- To handle powers, a common trick is to consider for x, y ≥ 0, z ∈ ℝ, ((x + y), (x - y, 2 * z)^T) ∈ SO, which is equivalent to z ^ 2 ≤ x * y. -/ -lemma soCone_add_sub_two_mul_of_nonneg {x y : ℝ} (z : ℝ) (hx : 0 ≤ x) (hy : 0 ≤ y) : +lemma soCone_add_sub_two_mul_of_nonneg {x y : ℝ} (z : ℝ) + (hx : 0 ≤ x) (hy : 0 ≤ y) : soCone (x + y) ![x - y, 2 * z] ↔ z ^ (2 : ℝ) ≤ x * y := by have hxy := add_nonneg hx hy conv => lhs; unfold soCone; simp [sqrt_le_left hxy, ←le_sub_iff_add_le'] diff --git a/CvxLean/Tactic/Convexify/Convexify.lean b/CvxLean/Tactic/Convexify/Convexify.lean index d6002104..9edbcd8b 100644 --- a/CvxLean/Tactic/Convexify/Convexify.lean +++ b/CvxLean/Tactic/Convexify/Convexify.lean @@ -19,19 +19,17 @@ def EggMinimization.ofOCTree (oc : OC (String × EggTree)) : /-- Given the rewrite name and direction from egg's output, find the appropriate tactic in the environment. It also returns a bool to indicate if the proof needs -an intermediate equality step. Otherwise, the tactic will be applied directly. -/ -def findTactic (isEquiv atObjFun : Bool) (rewriteName : String) (direction : EggRewriteDirection) : +an intermediate equality step. Otherwise, the tactic will be applied directly. +-/ +def findTactic (isEquiv atObjFun : Bool) (rewriteName : String) + (direction : EggRewriteDirection) : MetaM (Bool × TSyntax `tactic) := do match ← getTacticFromRewriteName rewriteName with | some (tac, mapObjFun) => if mapObjFun then if isEquiv then - -- NOTE: The proof obligations of our only two maps right now (log and - -- square) are solved by positivity. - return (true, ← `(tactic| positivity)) + return (true, ← `(tactic| arith)) else - -- NOTE: Only instance of a rewriting without proving an intermediate - -- equality. return (false, tac) else match direction with diff --git a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean index fdd6a74d..d284b767 100644 --- a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean +++ b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean @@ -51,6 +51,7 @@ def _root_.EggTree.opMap : HashMap String (String × Nat × Array String) := ("inv", ("inv", 1, #[])), ("abs", ("abs", 1, #[])), ("sqrt", ("sqrt", 1, #[])), + ("powOne", ("pow", 1, #["1"])), ("sq", ("pow", 1, #["2"])), ("powNegOne", ("pow", 1, #["-1"])), ("powNegTwo", ("pow", 1, #["-2"])), diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index 35eccdb7..e4ce6aed 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -208,10 +208,10 @@ register_rewrite_map "sqrt_eq_rpow-rev" ; "(pow ?a 0.5)" => "(sqrt ?a)" := simp_or_rw [←Real.sqrt_eq_rpow]; register_rewrite_map "pow_two"; "(pow ?a 2)" => "(mul ?a ?a)" := - simp_or_rw [pow_two (M := ℝ)]; + (norm_cast; simp_or_rw [pow_two]); register_rewrite_map "pow_two-rev"; "(mul ?a ?a)" => "(pow ?a 2)" := - simp_or_rw [←pow_two (M := ℝ)]; + (norm_cast; simp_or_rw [←pow_two]); register_rewrite_map "pow_half_two"; "(pow (pow ?a 0.5) 2)" => "?a" := simp_or_rw [Real.pow_half_two (by arith)]; diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean index 366327e2..c967b899 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean @@ -10,6 +10,16 @@ namespace CvxLean open Real +declare_atom powOne [affine] (x : ℝ)+ : x ^ (1 : ℝ) := +bconditions +homogenity by + simp +additivity by + simp +optimality by + intros _ h + simp [h] + declare_atom powNegOne [convex] (x : ℝ)- : x ^ (-1) := vconditions (hx : 0 < x) diff --git a/CvxLean/Test/Convexify/Unit.lean b/CvxLean/Test/Convexify/Unit.lean index e27b8307..df0e08c6 100644 --- a/CvxLean/Test/Convexify/Unit.lean +++ b/CvxLean/Test/Convexify/Unit.lean @@ -905,7 +905,7 @@ time_cmd equivalence mulPowRevConstrRed/mulPowRevConstrAuto : mulPowRevConstr := -- pow_mul-rev (obj) def powMulRevObj := optimization (x : ℝ) - minimize ((x ^ 2) ^ 2 : ℝ) + minimize ((x ^ 1) ^ 2 : ℝ) subject to hx : 0 < x @@ -973,8 +973,8 @@ def divPowRevConstr := hx : 0 < x h : ((x + x) ^ 2 / x ^ 2) ≤ x --- time_cmd equivalence divPowRevConstrRed/divPowRevConstrAuto : divPowRevConstr := by --- convexify +time_cmd equivalence divPowRevConstrRed/divPowRevConstrAuto : divPowRevConstr := by + convexify -- TODO(RFM): does not rewrite the correct term. -- #print divPowRevConstrAuto From b4c2e365a5f63b2a76596b9287a8b565c0127c0f Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 10:25:55 -0500 Subject: [PATCH 123/189] feat: make `arith` call `linarith` --- CvxLean/Tactic/Arith/Arith.lean | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/CvxLean/Tactic/Arith/Arith.lean b/CvxLean/Tactic/Arith/Arith.lean index fed6818f..f5f7f816 100644 --- a/CvxLean/Tactic/Arith/Arith.lean +++ b/CvxLean/Tactic/Arith/Arith.lean @@ -71,8 +71,14 @@ elab (name := prepare_positivity) "prepare_positivity" : tactic => do end Tactic +syntax "positivity!" : tactic + +macro_rules + | `(tactic| positivity!) => + `(tactic| cases_and; prepare_positivity; positivity) + syntax "arith" : tactic macro_rules | `(tactic| arith) => - `(tactic| cases_and; prepare_positivity; positivity) + `(tactic| (first | linarith | positivity!)) From 76eed4d4f06eb9f474fbce309d801013eaca9f57 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 10:58:09 -0500 Subject: [PATCH 124/189] feat: `norm_num_simp_pow` --- CvxLean/Tactic/Arith/Arith.lean | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/CvxLean/Tactic/Arith/Arith.lean b/CvxLean/Tactic/Arith/Arith.lean index f5f7f816..aa69c301 100644 --- a/CvxLean/Tactic/Arith/Arith.lean +++ b/CvxLean/Tactic/Arith/Arith.lean @@ -77,8 +77,14 @@ macro_rules | `(tactic| positivity!) => `(tactic| cases_and; prepare_positivity; positivity) +syntax "norm_num_simp_pow" : tactic + +macro_rules + | `(tactic| norm_num_simp_pow) => + `(tactic| norm_num [Real.rpow_neg]) + syntax "arith" : tactic macro_rules | `(tactic| arith) => - `(tactic| (first | linarith | positivity!)) + `(tactic| (first | linarith | positivity! | norm_num_simp_pow)) From eb43890a63973ddbd559dc981a419dc10ddd7b27 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 10:58:29 -0500 Subject: [PATCH 125/189] fix: `Vec.const` and `Vec.toMatrix` generalized --- CvxLean/Lib/Math/Data/Vec.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/CvxLean/Lib/Math/Data/Vec.lean b/CvxLean/Lib/Math/Data/Vec.lean index acde2c40..65a1ce0d 100644 --- a/CvxLean/Lib/Math/Data/Vec.lean +++ b/CvxLean/Lib/Math/Data/Vec.lean @@ -9,10 +9,10 @@ def abs [Abs α] (x : m → α) : m → α := instance [Abs α] : Abs (m → α) := ⟨abs⟩ -def const (n : ℕ) (k : ℝ) : Fin n → ℝ := +def const (n : ℕ) (k : α) : Fin n → α := fun _ => k -def toMatrix {n : ℕ} (x : Fin n → ℝ) : Fin n → Fin 1 → ℝ := +def toMatrix {n : ℕ} (x : Fin n → α) : Fin n → Fin 1 → α := fun i => ![x i] section AddCommMonoid From 1a1e824c214c4de09db81a729d0e3fe46b7379c5 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 11:18:39 -0500 Subject: [PATCH 126/189] feat: new atom `inv` --- CvxLean/Tactic/DCP/AtomLibrary/Fns/Inv.lean | 100 ++++++++++++++++++++ 1 file changed, 100 insertions(+) create mode 100644 CvxLean/Tactic/DCP/AtomLibrary/Fns/Inv.lean diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Inv.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Inv.lean new file mode 100644 index 00000000..1c14340f --- /dev/null +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Inv.lean @@ -0,0 +1,100 @@ +import CvxLean.Tactic.DCP.Atoms +import CvxLean.Tactic.DCP.AtomLibrary.Sets.Cones +import CvxLean.Tactic.DCP.AtomLibrary.Sets.Le +import CvxLean.Tactic.DCP.AtomLibrary.Fns.VecCons +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Add +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Sub +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Mul +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Power + +namespace CvxLean + +open Real + +-- TODO: These are just the same as `powNegOne`. We need them for pattern +-- mathcing. The issue is that implementationConstraints does not know about +-- vconditions. If we fix that, then we can just say (c : x ^ (-1) ≤ t). + +declare_atom inv [convex] (x : ℝ)- : x⁻¹ := +vconditions + (hx : 0 < x) +implementationVars (t : ℝ) +implementationObjective (t) +implementationConstraints + (c1 : soCone (t + x) ![t - x, 2]) + (c2 : 0 ≤ t) + (c3 : 0 ≤ x) +solution + (t := x⁻¹) +solutionEqualsAtom rfl +feasibility + (c1 : by + rw [Real.inv_eq_pow_neg_one hx] + exact powNegOne.feasibility0 x hx) + (c2 : by + rw [Real.inv_eq_pow_neg_one hx] + have hinv : 0 < x ^ (-1) := by rwa [rpow_neg_one, inv_pos] + exact le_of_lt hinv) + (c3 : le_of_lt hx) +optimality by + intros y hy + rw [soCone_add_sub_two_of_nonneg c2 c3] at c1 + have hxpos : 0 < x := by + cases (lt_or_eq_of_le c3) with + | inl h => exact h + | inr h => rw [←h] at c1; linarith + have hypos := lt_of_lt_of_le hxpos hy + rw [inv_eq_pow_neg_one hypos] + rw [rpow_neg_one, inv_eq_one_div, div_le_iff hypos] + apply le_trans c1 + apply mul_le_mul_of_nonneg_left hy c2 +vconditionElimination + (hx : by + intros y hy + rw [soCone_add_sub_two_of_nonneg c2 c3] at c1 + cases (lt_or_eq_of_le c3) with + | inl h => linarith + | inr h => rw [←h] at c1; linarith) + +declare_atom oneDiv [convex] (x : ℝ)- : 1 / x := +vconditions + (hx : 0 < x) +implementationVars (t : ℝ) +implementationObjective (t) +implementationConstraints + (c1 : soCone (t + x) ![t - x, 2]) + (c2 : 0 ≤ t) + (c3 : 0 ≤ x) +solution + (t := 1 / x) +solutionEqualsAtom rfl +feasibility + (c1 : by + rw [Real.one_div_eq_pow_neg_one hx] + exact powNegOne.feasibility0 x hx) + (c2 : by + rw [Real.one_div_eq_pow_neg_one hx] + have hinv : 0 < x ^ (-1) := by rwa [rpow_neg_one, inv_pos] + exact le_of_lt hinv) + (c3 : le_of_lt hx) +optimality by + intros y hy + rw [soCone_add_sub_two_of_nonneg c2 c3] at c1 + have hxpos : 0 < x := by + cases (lt_or_eq_of_le c3) with + | inl h => exact h + | inr h => rw [←h] at c1; linarith + have hypos := lt_of_lt_of_le hxpos hy + rw [one_div_eq_pow_neg_one hypos] + rw [rpow_neg_one, inv_eq_one_div, div_le_iff hypos] + apply le_trans c1 + apply mul_le_mul_of_nonneg_left hy c2 +vconditionElimination + (hx : by + intros y hy + rw [soCone_add_sub_two_of_nonneg c2 c3] at c1 + cases (lt_or_eq_of_le c3) with + | inl h => linarith + | inr h => rw [←h] at c1; linarith) + +end CvxLean From 4523ec70a3c73ebd152cd4880d788259ec664edf Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 11:19:22 -0500 Subject: [PATCH 127/189] chore: simplify library lemmas about Real --- CvxLean/Lib/Math/Data/Real.lean | 28 ++++++++++++++-------------- 1 file changed, 14 insertions(+), 14 deletions(-) diff --git a/CvxLean/Lib/Math/Data/Real.lean b/CvxLean/Lib/Math/Data/Real.lean index 766906f0..7e60c83f 100644 --- a/CvxLean/Lib/Math/Data/Real.lean +++ b/CvxLean/Lib/Math/Data/Real.lean @@ -8,7 +8,7 @@ namespace Real section Functions noncomputable def entr (x : Real) := - -(x * Real.log x) + -(x * log x) noncomputable def huber (x : Real) := if abs x ≤ 1 then x ^ 2 else 2 * abs x - 1 @@ -39,43 +39,43 @@ end Instances /- Lemmas used in `RewriteMapLibrary`. -/ section Lemmas -lemma Real.log_eq_log {x y : ℝ} (hx : 0 < x) (hy : 0 < y) : - Real.log x = Real.log y ↔ x = y := +lemma log_eq_log {x y : ℝ} (hx : 0 < x) (hy : 0 < y) : + log x = log y ↔ x = y := ⟨fun h => by have hxmem := Set.mem_Ioi.2 hx have hymem := Set.mem_Ioi.2 hy have heq : Set.restrict (Set.Ioi 0) log ⟨x, hxmem⟩ = Set.restrict (Set.Ioi 0) log ⟨y, hymem⟩ := by simp [h] - have h := Real.log_injOn_pos.injective heq + have h := log_injOn_pos.injective heq simp [Subtype.eq] at h exact h, fun h => by rw [h]⟩ -lemma Real.div_pow_eq_mul_pow_neg {a b c : ℝ} (hb : 0 ≤ b) : +lemma div_pow_eq_mul_pow_neg {a b c : ℝ} (hb : 0 ≤ b) : a / (b ^ c) = a * b ^ (-c) := by rw [div_eq_mul_inv, ←rpow_neg hb] -lemma Real.one_div_eq_pow_neg_one {a : ℝ} (ha : 0 < a) : 1 / a = a ^ (-1) := by - rw [Real.rpow_neg (le_of_lt ha), rpow_one, div_eq_mul_inv, one_mul] +lemma one_div_eq_pow_neg_one {a : ℝ} (ha : 0 < a) : 1 / a = a ^ (-1) := by + rw [rpow_neg (le_of_lt ha), rpow_one, div_eq_mul_inv, one_mul] -lemma Real.inv_eq_pow_neg_one {a : ℝ} (ha : 0 < a) : a⁻¹ = a ^ (-1) := by - rw [inv_eq_one_div, Real.one_div_eq_pow_neg_one ha] +lemma inv_eq_pow_neg_one {a : ℝ} (ha : 0 < a) : a⁻¹ = a ^ (-1) := by + rw [inv_eq_one_div, one_div_eq_pow_neg_one ha] -lemma Real.pow_half_two {x : ℝ} (hx : 0 ≤ x) : (x ^ (1 / 2)) ^ 2 = x := by - show Real.rpow (Real.rpow _ _) _ = _ +lemma pow_half_two {x : ℝ} (hx : 0 ≤ x) : (x ^ (1 / 2)) ^ 2 = x := by + show rpow (rpow _ _) _ = _ rw [rpow_eq_pow, rpow_eq_pow, ← rpow_mul hx] norm_num -lemma Real.pow_two_le_pow_two {x y : ℝ} (hx : 0 ≤ x) (hy : 0 ≤ y) : +lemma pow_two_le_pow_two {x y : ℝ} (hx : 0 ≤ x) (hy : 0 ≤ y) : x ^ 2 ≤ y ^ 2 ↔ x ≤ y := by rw [rpow_two, rpow_two, sq_le_sq, abs_of_nonneg hx, abs_of_nonneg hy] -lemma Real.binomial_two (x y : ℝ) : +lemma binomial_two (x y : ℝ) : (x + y) ^ 2 = x ^ 2 + (2 * (x * y) + y ^ 2) := by simp only [rpow_two]; ring -lemma Real.exp_neg_eq_one_div (x : ℝ) : exp (-x) = 1 / exp x := by +lemma exp_neg_eq_one_div (x : ℝ) : exp (-x) = 1 / exp x := by rw [exp_neg, inv_eq_one_div] end Lemmas From 25f6044fa75c0c00d48391153bc2d01f5b126dcf Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 11:19:39 -0500 Subject: [PATCH 128/189] feat: only use `arith` in `dcp` --- CvxLean/Tactic/DCP/Dcp.lean | 8 +++----- 1 file changed, 3 insertions(+), 5 deletions(-) diff --git a/CvxLean/Tactic/DCP/Dcp.lean b/CvxLean/Tactic/DCP/Dcp.lean index d5d1c195..e533fcbd 100644 --- a/CvxLean/Tactic/DCP/Dcp.lean +++ b/CvxLean/Tactic/DCP/Dcp.lean @@ -1,6 +1,4 @@ import Mathlib.Tactic.NormNum -import Mathlib.Tactic.Positivity -import Mathlib.Tactic.Linarith import CvxLean.Tactic.DCP.AtomExt import CvxLean.Syntax.Minimization import CvxLean.Meta.Util.Meta @@ -128,9 +126,9 @@ where | some fvarId => bconds := bconds.push (mkFVar fvarId) | none => - -- Try to prove simple bconditions by norm_num. + -- Try to prove simple bconditions by arith. let (e, _) ← Lean.Elab.Term.TermElabM.run $ Lean.Elab.Term.commitIfNoErrors? $ do - let v ← Lean.Elab.Term.elabTerm (← `(by norm_num)).raw (some bcondType) + let v ← Lean.Elab.Term.elabTerm (← `(by arith)).raw (some bcondType) Lean.Elab.Term.synthesizeSyntheticMVarsNoPostponing let v ← instantiateMVars v return v @@ -263,7 +261,7 @@ partial def findVConditions (originalConstrVars : Array LocalDecl) (constraints let vcondProofTy ← mkForallFVars args vcondProofTyBody let (e, _) ← Lean.Elab.Term.TermElabM.run <| Lean.Elab.Term.commitIfNoErrors? <| do - let tac ← `(by intros; try { arith <;> linarith <;> norm_num }) + let tac ← `(by intros; try { arith }) let v ← Lean.Elab.Term.elabTerm tac.raw (some vcondProofTy) Lean.Elab.Term.synthesizeSyntheticMVarsNoPostponing instantiateMVars v From 63d72a2c523730a04da3a93362db4c75346b5046 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 11:37:49 -0500 Subject: [PATCH 129/189] feat: introduce `OpArgTag` to deal with 1/x in `opMap` --- .../Convexify/Egg/FromMinimization.lean | 99 +++++++++++-------- 1 file changed, 59 insertions(+), 40 deletions(-) diff --git a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean index d284b767..2e828a3d 100644 --- a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean +++ b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean @@ -31,45 +31,50 @@ partial def _root_.EggTree.surroundVars (t : Tree String String) (vars : List St | Tree.leaf s => if s ∈ vars then Tree.node "var" #[Tree.leaf s] else t | Tree.node n children => Tree.node n (children.map (surroundVars · vars)) +inductive _root_.EggTree.OpArgTag + | arg + | val (s : String) + /-- Mapping between atom names that come from the unchecked tree construction and their egg counterparts. `prob`, `objFun`, `constr` and `constrs` are special cases. The rest of names come from the atom library. It also returns the arity of the operation and some extra arguments. -/ -def _root_.EggTree.opMap : HashMap String (String × Nat × Array String) := +def _root_.EggTree.opMap : HashMap String (String × Array EggTree.OpArgTag) := HashMap.ofList [ - ("prob", ("prob", 2, #[])), - ("objFun", ("objFun", 1, #[])), - ("constr", ("constr", 2, #[])), - ("constrs", ("constrs", 0, #[])), - ("maximizeNeg", ("neg", 1, #[])), - ("var", ("var", 1, #[])), - ("param", ("param", 1, #[])), - ("eq", ("eq", 2, #[])), - ("le", ("le", 2, #[])), - ("lt", ("lt", 2, #[])), - ("neg", ("neg", 1, #[])), - ("inv", ("inv", 1, #[])), - ("abs", ("abs", 1, #[])), - ("sqrt", ("sqrt", 1, #[])), - ("powOne", ("pow", 1, #["1"])), - ("sq", ("pow", 1, #["2"])), - ("powNegOne", ("pow", 1, #["-1"])), - ("powNegTwo", ("pow", 1, #["-2"])), - ("log", ("log", 1, #[])), - ("exp", ("exp", 1, #[])), - ("xexp", ("xexp", 1, #[])), - ("entr", ("entr", 1, #[])), - ("min", ("min", 2, #[])), - ("max", ("max", 2, #[])), - ("add", ("add", 2, #[])), - ("sub", ("sub", 2, #[])), - ("mul1", ("mul", 2, #[])), - ("mul2", ("mul", 2, #[])), - ("div", ("div", 2, #[])), - ("quadOverLin", ("qol", 2, #[])), - ("geoMean", ("geo", 2, #[])), - ("logSumExp", ("lse", 2, #[])), - ("norm2₂", ("norm2", 2, #[])) + ("prob", ("prob", #[.arg, .arg])), + ("objFun", ("objFun", #[.arg])), + ("constr", ("constr", #[.arg, .arg])), + ("constrs", ("constrs", #[])), + ("maximizeNeg", ("neg", #[.arg])), + ("var", ("var", #[.arg])), + ("param", ("param", #[.arg])), + ("eq", ("eq", #[.arg])), + ("le", ("le", #[.arg, .arg])), + ("lt", ("lt", #[.arg, .arg])), + ("neg", ("neg", #[.arg])), + ("inv", ("inv", #[.arg])), + ("oneDiv", ("div", #[.val "1", .arg])), + ("abs", ("abs", #[.arg])), + ("sqrt", ("sqrt", #[.arg])), + ("powOne", ("pow", #[.arg, .val "1"])), + ("sq", ("pow", #[.arg, .val "2"])), + ("powNegOne", ("pow", #[.arg, .val "-1"])), + ("powNegTwo", ("pow", #[.arg, .val "-2"])), + ("log", ("log", #[.arg])), + ("exp", ("exp", #[.arg])), + ("xexp", ("xexp", #[.arg])), + ("entr", ("entr", #[.arg])), + ("min", ("min", #[.arg, .arg])), + ("max", ("max", #[.arg, .arg])), + ("add", ("add", #[.arg, .arg])), + ("sub", ("sub", #[.arg, .arg])), + ("mul1", ("mul", #[.arg, .arg])), + ("mul2", ("mul", #[.arg, .arg])), + ("div", ("div", #[.arg, .arg])), + ("quadOverLin", ("qol", #[.arg, .arg])), + ("geoMean", ("geo", #[.arg, .arg])), + ("logSumExp", ("lse", #[.arg, .arg])), + ("norm2₂", ("norm2", #[.arg, .arg])) ] /-- Traverse the tree and use `EggTree.opMap` to align the names of the @@ -78,12 +83,26 @@ partial def _root_.EggTree.adjustOps (t : Tree String String) : MetaM (Tree String String) := do match t with | Tree.node op children => - if let some (op', arity, extraArgs) := EggTree.opMap.find? op then - if children.size ≠ arity && op' != "constrs" then - throwError s!"The operator {op} has arity {children.size}, but it should have arity {arity}." - let children' ← children.mapM adjustOps - let children' := children' ++ extraArgs.map Tree.leaf - return Tree.node op' children' + if let some (newOp, argTags) := EggTree.opMap.find? op then + let mut children ← children.mapM adjustOps + let mut newChildren := #[] + let mut arityError := false + for argTag in argTags do + match argTag with + | .arg => + if children.size > 0 then + newChildren := newChildren.push children[0]! + children := children.drop 1 + else + arityError := true + break + | .val s => + newChildren := newChildren.push (Tree.leaf s) + arityError := arityError || children.size > 0 + if arityError then + throwError s!"The op {op} was passed {children.size} arguments." + else + return Tree.node newOp newChildren else throwError s!"The atom {op} is not supported by the `convexify` tactic." | Tree.leaf "unknown" => throwError "Unknown atom." From 7c97bbce127b67dd029703c4d73ca8ec395e81f6 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 11:46:01 -0500 Subject: [PATCH 130/189] feat: add `inv` to library --- CvxLean/Tactic/DCP/AtomLibrary/Fns/All.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/All.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/All.lean index 49b4c896..2994a019 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/All.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/All.lean @@ -9,6 +9,7 @@ import CvxLean.Tactic.DCP.AtomLibrary.Fns.Exp import CvxLean.Tactic.DCP.AtomLibrary.Fns.FromBlocks import CvxLean.Tactic.DCP.AtomLibrary.Fns.GeoMean import CvxLean.Tactic.DCP.AtomLibrary.Fns.Huber +import CvxLean.Tactic.DCP.AtomLibrary.Fns.Inv import CvxLean.Tactic.DCP.AtomLibrary.Fns.KLDiv import CvxLean.Tactic.DCP.AtomLibrary.Fns.Log import CvxLean.Tactic.DCP.AtomLibrary.Fns.LogDet From fc5acc347bace466b4998fad013981173081f3b1 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 11:49:05 -0500 Subject: [PATCH 131/189] fix: number of arguments in `eq` --- CvxLean/Tactic/Convexify/Egg/FromMinimization.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean index 2e828a3d..cc0a6bb9 100644 --- a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean +++ b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean @@ -48,7 +48,7 @@ def _root_.EggTree.opMap : HashMap String (String × Array EggTree.OpArgTag) := ("maximizeNeg", ("neg", #[.arg])), ("var", ("var", #[.arg])), ("param", ("param", #[.arg])), - ("eq", ("eq", #[.arg])), + ("eq", ("eq", #[.arg, .arg])), ("le", ("le", #[.arg, .arg])), ("lt", ("lt", #[.arg, .arg])), ("neg", ("neg", #[.arg])), @@ -85,6 +85,7 @@ partial def _root_.EggTree.adjustOps (t : Tree String String) : | Tree.node op children => if let some (newOp, argTags) := EggTree.opMap.find? op then let mut children ← children.mapM adjustOps + let argsGiven := children.size let mut newChildren := #[] let mut arityError := false for argTag in argTags do @@ -100,7 +101,7 @@ partial def _root_.EggTree.adjustOps (t : Tree String String) : newChildren := newChildren.push (Tree.leaf s) arityError := arityError || children.size > 0 if arityError then - throwError s!"The op {op} was passed {children.size} arguments." + throwError s!"The op {op} was passed {argsGiven} arguments." else return Tree.node newOp newChildren else From 490909c516ee2935eb5d1e4882e02529c81d8115 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 12:25:50 -0500 Subject: [PATCH 132/189] fix: do not use `arith` for rewrite maps --- .../Tactic/Convexify/RewriteMapLibrary.lean | 68 +++++++++---------- 1 file changed, 34 insertions(+), 34 deletions(-) diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index e4ce6aed..37076ba7 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -30,7 +30,7 @@ register_objFun_rewrite_map "map_objFun_sq"; "(prob (objFun ?a) ?cs)" => "(prob /- Equality rules. -/ register_rewrite_map "log_eq_log" ; "(eq ?a ?b)" => "(eq (log ?a) (log ?b))" := - rewrite [Real.log_eq_log (by arith) (by arith)]; + rewrite [Real.log_eq_log (by positivity!) (by positivity!)]; /- Less than or equal rules. -/ @@ -48,25 +48,25 @@ register_rewrite_map "sub_le_iff_le_add-rev"; "(le ?a (add ?b ?c))" => "(le (sub rewrite [←sub_le_iff_le_add]; register_rewrite_map "div_le_iff" ; "(le (div ?a ?b) ?c)" => "(le ?a (mul ?b ?c))" := - rewrite [div_le_iff (by arith)]; + rewrite [div_le_iff (by positivity!)]; register_rewrite_map "div_le_iff-rev" ; "(le (div ?a ?b) ?c)" => "(le ?a (mul ?b ?c))" := - rewrite [←div_le_iff (by arith)]; + rewrite [←div_le_iff (by positivity!)]; register_rewrite_map "div_le_one-rev" ; "(le ?a ?b)" => "(le (div ?a ?b) 1)" := - rewrite [←div_le_one (by arith)]; + rewrite [←div_le_one (by positivity!)]; register_rewrite_map "log_le_log" ; "(le (log ?a) (log ?b))" => "(le ?a ?b)" := - rewrite [Real.log_le_log (by arith) (by arith)]; + rewrite [Real.log_le_log (by positivity!) (by positivity!)]; register_rewrite_map "log_le_log-rev" ; "(le ?a ?b)" => "(le (log ?a) (log ?b))" := - rewrite [←Real.log_le_log (by arith) (by arith)]; + rewrite [←Real.log_le_log (by positivity!) (by positivity!)]; register_rewrite_map "pow_two_le_pow_two"; "(le (pow ?a 2) (pow ?b 2))" => "(le ?a ?b)":= - rewrite [Real.pow_two_le_pow_two (by arith) (by arith)]; + rewrite [Real.pow_two_le_pow_two (by positivity!) (by positivity!)]; register_rewrite_map "pow_two_le_pow_two-rev"; "(le ?a ?b)" => "(le (pow ?a 2) (pow ?b 2))" := - rewrite [←Real.pow_two_le_pow_two (by arith) (by arith)]; + rewrite [←Real.pow_two_le_pow_two (by positivity!) (by positivity!)]; /- Field rules. -/ @@ -145,7 +145,7 @@ register_rewrite_map "mul_div-rev" ; "(div (mul ?a ?b) ?c)" => "(mul ?a (div ?b simp_or_rw [←mul_div (G := ℝ)]; register_rewrite_map "div_self" ; "(div ?a ?a)" => "1" := - simp_or_rw [div_self (G₀ := ℝ) (by arith)]; + simp_or_rw [div_self (G₀ := ℝ) (by positivity!)]; register_rewrite_map "inv_eq_one_div" ; "(inv ?a)" => "(div 1 ?a)" := simp_or_rw [inv_eq_one_div (G := ℝ)]; @@ -163,43 +163,43 @@ register_rewrite_map "pow_one"; "(pow ?a 1)" => "?a" := simp_or_rw [Real.rpow_one]; register_rewrite_map "pow_add"; "(pow ?a (add ?b ?c))" => "(mul (pow ?a ?b) (pow ?a ?c))" := - simp_or_rw [Real.rpow_add (by arith)]; + simp_or_rw [Real.rpow_add (by positivity!)]; register_rewrite_map "pow_add-rev"; "(mul (pow ?a ?b) (pow ?a ?c))" => "(pow ?a (add ?b ?c))" := - simp_or_rw [←Real.rpow_add (by arith)]; + simp_or_rw [←Real.rpow_add (by positivity!)]; register_rewrite_map "mul_pow"; "(mul (pow ?a ?n) (pow ?b ?n))" => "(pow (mul ?a ?b) ?n)" := - simp_or_rw [Real.mul_rpow (by arith) (by arith)]; + simp_or_rw [Real.mul_rpow (by positivity!) (by positivity!)]; register_rewrite_map "mul_pow-rev"; "(mul (pow ?a ?n) (pow ?b ?n))" => "(pow (mul ?a ?b) ?n)" := - simp_or_rw [←Real.mul_rpow (by arith) (by arith)]; + simp_or_rw [←Real.mul_rpow (by positivity!) (by positivity!)]; register_rewrite_map "pow_mul"; "(pow ?a (mul ?n ?m))" => "(pow (pow ?a ?n) ?m)" := - simp_or_rw [Real.rpow_mul (by arith)]; + simp_or_rw [Real.rpow_mul (by positivity!)]; register_rewrite_map "pow_mul-rev"; "(pow (pow ?a ?n) ?m)" => "(pow ?a (mul ?n ?m))" := - simp_or_rw [←Real.rpow_mul (by arith)]; + simp_or_rw [←Real.rpow_mul (by positivity!)]; register_rewrite_map "div_pow"; "(pow (div ?a ?b) ?n)" => "(div (pow ?a ?n) (pow ?b ?n))" := - simp_or_rw [Real.div_rpow (by arith) (by arith)]; + simp_or_rw [Real.div_rpow (by positivity!) (by positivity!)]; register_rewrite_map "div_pow-rev"; "(div (pow ?a ?n) (pow ?b ?n))" => "(pow (div ?a ?b) ?n)" := - simp_or_rw [←Real.div_rpow (by arith) (by arith)]; + simp_or_rw [←Real.div_rpow (by positivity!) (by positivity!)]; register_rewrite_map "pow_sub"; "(pow ?a (sub ?b ?c))" => "(div (pow ?a ?b) (pow ?a ?c))" := - simp_or_rw [Real.rpow_sub (by arith)]; + simp_or_rw [Real.rpow_sub (by positivity!)]; register_rewrite_map "pow_sub-rev"; "(div (pow ?a ?b) (pow ?a ?c))" => "(pow ?a (sub ?b ?c))" := - simp_or_rw [←Real.rpow_sub (by arith)]; + simp_or_rw [←Real.rpow_sub (by positivity!)]; register_rewrite_map "div_pow_eq_mul_pow_neg" ; "(div ?a (pow ?b ?c))" => "(mul ?a (pow ?b (neg ?c)))" := - simp_or_rw [Real.div_pow_eq_mul_pow_neg (by arith)]; + simp_or_rw [Real.div_pow_eq_mul_pow_neg (by positivity!)]; register_rewrite_map "div_pow_eq_mul_pow_neg-rev" ; "(div ?a (pow ?b ?c))" => "(mul ?a (pow ?b (neg ?c)))" := - simp_or_rw [←Real.div_pow_eq_mul_pow_neg (by arith)]; + simp_or_rw [←Real.div_pow_eq_mul_pow_neg (by positivity!)]; register_rewrite_map "one_div_eq_pow_neg_one"; "(div 1 ?a)" => "(pow ?a (neg 1))" := - simp_or_rw [Real.one_div_eq_pow_neg_one (by arith)]; + simp_or_rw [Real.one_div_eq_pow_neg_one (by positivity!)]; register_rewrite_map "sqrt_eq_rpow" ; "(sqrt ?a)" => "(pow ?a 0.5)" := simp_or_rw [Real.sqrt_eq_rpow]; @@ -214,17 +214,17 @@ register_rewrite_map "pow_two-rev"; "(mul ?a ?a)" => "(pow ?a 2)" := (norm_cast; simp_or_rw [←pow_two]); register_rewrite_map "pow_half_two"; "(pow (pow ?a 0.5) 2)" => "?a" := - simp_or_rw [Real.pow_half_two (by arith)]; + simp_or_rw [Real.pow_half_two (by positivity!)]; -- TODO(RFM): Technically ← but no pattern to match on otherwise. register_rewrite_map "pow_half_two-rev"; "?a" => "(pow (pow ?a 0.5) 2)" := - simp_or_rw [Real.pow_half_two (by arith)]; + simp_or_rw [Real.pow_half_two (by positivity!)]; register_rewrite_map "binomial_two"; "(pow (add ?a ?b) 2)" => "(add (pow ?a 2) (add (mul 2 (mul ?a ?b)) (pow ?b 2)))" := simp_or_rw [Real.binomial_two]; register_rewrite_map "inv_eq_pow_neg_one"; "(inv ?a)" => "(pow ?a (neg 1))" := - simp_or_rw [Real.inv_eq_pow_neg_one (by arith)]; + simp_or_rw [Real.inv_eq_pow_neg_one (by positivity!)]; /- Exponential and logarithm rules. -/ @@ -254,22 +254,22 @@ register_rewrite_map "exp_neg_eq_one_div-rev" ; "(div 1 (exp ?a))" => "(exp (neg simp_or_rw [←Real.exp_neg_eq_one_div]; register_rewrite_map "log_mul" ; "(log (mul ?a ?b))" => "(add (log ?a) (log ?b))" := - simp_or_rw [Real.log_mul (by arith) (by arith)]; + simp_or_rw [Real.log_mul (by positivity!) (by positivity!)]; register_rewrite_map "log_mul-rev"; "(add (log ?a) (log ?b))" => "(log (mul ?a ?b))" := - simp_or_rw [←Real.log_mul (by arith) (by arith)]; + simp_or_rw [←Real.log_mul (by positivity!) (by positivity!)]; register_rewrite_map "log_div" ; "(log (div ?a ?b))" => "(sub (log ?a) (log ?b))" := - simp_or_rw [Real.log_div (by arith) (by arith)]; + simp_or_rw [Real.log_div (by positivity!) (by positivity!)]; register_rewrite_map "log_div-rev"; "(sub (log ?a) (log ?b))" => "(log (div ?a ?b))" := - simp_or_rw [←Real.log_div (by arith) (by arith)]; + simp_or_rw [←Real.log_div (by positivity!) (by positivity!)]; register_rewrite_map "log_pow" ; "(log (pow ?a ?b))" => "(mul ?b (log ?a))" := - simp_or_rw [Real.log_rpow (by arith)]; + simp_or_rw [Real.log_rpow (by positivity!)]; register_rewrite_map "log_pow-rev" ; "(mul ?b (log ?a))" => "(log (pow ?a ?b))" := - simp_or_rw [←Real.log_rpow (by arith)]; + simp_or_rw [←Real.log_rpow (by positivity!)]; -- NOTE: Special rule that only works if the exponent is a natural number. register_rewrite_map "log_pow_nat" ; "(log (pow ?a ?b))" => "(mul ?b (log ?a))" := @@ -279,7 +279,7 @@ register_rewrite_map "log_pow_nat-rev" ; "(mul ?b (log ?a))" => "(log (pow ?a ?b (norm_cast; simp_or_rw [←Real.log_pow]); register_rewrite_map "exp_log" ; "(exp (log ?a))" => "?a" := - simp_or_rw [Real.exp_log (by arith)]; + simp_or_rw [Real.exp_log (by positivity!)]; register_rewrite_map "log_exp" ; "(log (exp ?a))" => "?a" := simp_or_rw [Real.log_exp]; @@ -288,10 +288,10 @@ register_rewrite_map "log_exp" ; "(log (exp ?a))" => "?a" := /- Absolute value rules. -/ register_rewrite_map "abs_nonneg" ; "(abs ?a)" => "?a" := - simp_or_rw [abs_eq_self.mpr (by arith)]; + simp_or_rw [abs_eq_self.mpr (by positivity!)]; register_rewrite_map "abs_nonpos" ; "(abs ?a)" => "(neg ?a)" := - simp_or_rw [abs_eq_neg_self.mpr (by arith)]; + simp_or_rw [abs_eq_neg_self.mpr (by linarith)]; /- Atom folding. -/ From c5089bcead982f9cb97e92aa5218db3b6ac80078 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 12:29:46 -0500 Subject: [PATCH 133/189] test: recover unit tests --- CvxLean/Test/Convexify/Unit.lean | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/CvxLean/Test/Convexify/Unit.lean b/CvxLean/Test/Convexify/Unit.lean index df0e08c6..1219e1a8 100644 --- a/CvxLean/Test/Convexify/Unit.lean +++ b/CvxLean/Test/Convexify/Unit.lean @@ -959,8 +959,8 @@ def divPowRevObj := subject to hx : 0 < x --- time_cmd equivalence divPowRevObjRed/divPowRevObjAuto : divPowRevObj := by - -- convexify +time_cmd equivalence divPowRevObjRed/divPowRevObjAuto : divPowRevObj := by + convexify -- TODO(RFM): does not rewrite the correct term. -- #print divPowRevObjAuto @@ -975,7 +975,6 @@ def divPowRevConstr := time_cmd equivalence divPowRevConstrRed/divPowRevConstrAuto : divPowRevConstr := by convexify --- TODO(RFM): does not rewrite the correct term. -- #print divPowRevConstrAuto From 35f147ce904a960c80a8998c79efb9df4d5e8d21 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 12:36:17 -0500 Subject: [PATCH 134/189] feat: simplify powers in `convexify` and better error messages --- CvxLean/Tactic/Convexify/Convexify.lean | 13 +++++++++---- 1 file changed, 9 insertions(+), 4 deletions(-) diff --git a/CvxLean/Tactic/Convexify/Convexify.lean b/CvxLean/Tactic/Convexify/Convexify.lean index 9edbcd8b..61effd8d 100644 --- a/CvxLean/Tactic/Convexify/Convexify.lean +++ b/CvxLean/Tactic/Convexify/Convexify.lean @@ -108,7 +108,7 @@ output the remaining goals. -/ def evalStep (g : MVarId) (step : EggRewrite) (vars : List Name) (fvars : Array Expr) (domain : Expr) (numConstrTags : ℕ) (tagsMap : HashMap String ℕ) (isEquiv : Bool) : - TacticM (List MVarId) := do + TacticM (List MVarId) := withMainContext <| do let tag := step.location let tagNum := tagsMap.find! tag let atObjFun := tagNum == 0 @@ -207,9 +207,9 @@ def evalStep (g : MVarId) (step : EggRewrite) -- of approaches using `norm_num` in combination with the tactic provided -- by the user for this particular rewrite. let fullTac : Syntax ← `(tactic| intros; - try { norm_num_clean_up; $tacStx <;> norm_num } <;> - try { $tacStx <;> norm_num } <;> - try { norm_num }) + try { norm_num_clean_up; $tacStx <;> norm_num_simp_pow } <;> + try { $tacStx <;> norm_num_simp_pow } <;> + try { norm_num_simp_pow }) let gsAfterRw ← evalTacticAt fullTac gToRw if gsAfterRw.length == 0 then @@ -306,6 +306,11 @@ def evalConvexify : Tactic := fun stx => match stx with g := gs[0]! replaceMainGoal [g] + let gs ← getUnsolvedGoals + if gs.length != 1 then + replaceMainGoal gs + throwError "Failed to close subgoals." + normNumCleanUp (useSimp := false) saveTacticInfoForToken stx From d6996075f84d4762aa694aa0cdb4efdaf55dc06a Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 12:39:45 -0500 Subject: [PATCH 135/189] refactor: move `norm_num` variants --- CvxLean/Tactic/Arith/Arith.lean | 7 +----- CvxLean/Tactic/Arith/NormNumVariants.lean | 29 +++++++++++++++++++++++ CvxLean/Tactic/Convexify/Convexify.lean | 15 +----------- 3 files changed, 31 insertions(+), 20 deletions(-) create mode 100644 CvxLean/Tactic/Arith/NormNumVariants.lean diff --git a/CvxLean/Tactic/Arith/Arith.lean b/CvxLean/Tactic/Arith/Arith.lean index aa69c301..b9138c38 100644 --- a/CvxLean/Tactic/Arith/Arith.lean +++ b/CvxLean/Tactic/Arith/Arith.lean @@ -2,6 +2,7 @@ import Lean import Qq import Mathlib.Data.Real.Basic import Mathlib.Tactic.Positivity +import CvxLean.Tactic.Arith.NormNumVariants import CvxLean.Tactic.Arith.PositivityExt namespace Tactic @@ -77,12 +78,6 @@ macro_rules | `(tactic| positivity!) => `(tactic| cases_and; prepare_positivity; positivity) -syntax "norm_num_simp_pow" : tactic - -macro_rules - | `(tactic| norm_num_simp_pow) => - `(tactic| norm_num [Real.rpow_neg]) - syntax "arith" : tactic macro_rules diff --git a/CvxLean/Tactic/Arith/NormNumVariants.lean b/CvxLean/Tactic/Arith/NormNumVariants.lean new file mode 100644 index 00000000..3b7b91b4 --- /dev/null +++ b/CvxLean/Tactic/Arith/NormNumVariants.lean @@ -0,0 +1,29 @@ +import Lean +import Mathlib.Tactic.NormNum + +namespace CvxLean + +open Lean Elab Meta Tactic + +/-- Variant of `norm_num` used to get rid of the `OfScientific`s. -/ +def normNumCleanUp (useSimp : Bool) : TacticM Unit := + Mathlib.Meta.NormNum.elabNormNum mkNullNode mkNullNode (useSimp := useSimp) + +syntax (name := norm_num_clean_up) "norm_num_clean_up" : tactic + +@[tactic norm_num_clean_up] +def evalNormNumCleanUp : Tactic := fun stx => match stx with + | `(tactic| norm_num_clean_up) => do + normNumCleanUp (useSimp := false) + saveTacticInfoForToken stx + | _ => throwUnsupportedSyntax + + +/-- Variant of `norm_num` that simplifies powers. -/ +syntax "norm_num_simp_pow" : tactic + +macro_rules + | `(tactic| norm_num_simp_pow) => + `(tactic| norm_num [Real.rpow_neg]) + +end CvxLean diff --git a/CvxLean/Tactic/Convexify/Convexify.lean b/CvxLean/Tactic/Convexify/Convexify.lean index 61effd8d..c39e16c5 100644 --- a/CvxLean/Tactic/Convexify/Convexify.lean +++ b/CvxLean/Tactic/Convexify/Convexify.lean @@ -1,7 +1,7 @@ import Lean -import Mathlib.Tactic.NormNum import CvxLean.Meta.Minimization import CvxLean.Lib.Equivalence +import CvxLean.Tactic.Arith.NormNumVariants import CvxLean.Tactic.Convexify.RewriteMapExt import CvxLean.Tactic.Convexify.RewriteMapLibrary import CvxLean.Tactic.Convexify.Egg.All @@ -89,19 +89,6 @@ def rewriteWrapperApplyExpr (givenRange : Bool) (rwName : Name) (numArgs : Nat) let args ← Array.range numArgs |>.mapM fun _ => mkFreshExprMVar none return mkAppN (mkConst rwName) (signature ++ args ++ #[expected]) -/-- Version of `norm_num` used to get rid of the `OfScientific`s. -/ -def normNumCleanUp (useSimp : Bool) : TacticM Unit := - Mathlib.Meta.NormNum.elabNormNum mkNullNode mkNullNode (useSimp := useSimp) - -syntax (name := norm_num_clean_up) "norm_num_clean_up" : tactic - -@[tactic norm_num_clean_up] -def evalNormNumCleanUp : Tactic := fun stx => match stx with - | `(tactic| norm_num_clean_up) => do - normNumCleanUp (useSimp := false) - saveTacticInfoForToken stx - | _ => throwUnsupportedSyntax - /-- Given an egg rewrite and a current goal with all the necessary information about the minimization problem, we find the appropriate rewrite to apply, and output the remaining goals. -/ From 5b3b55851356f728a927808111eaff81030639cf Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 17:21:05 -0500 Subject: [PATCH 136/189] feat: grouping cones correctly --- CvxLean/Command/Solve.lean | 5 +- CvxLean/Tactic/Solver/Conic.lean | 72 +++++++++---------------- CvxLean/Tactic/Solver/Float/Coeffs.lean | 33 +++++++++--- CvxLean/Tactic/Solver/Mosek/CBF.lean | 3 +- 4 files changed, 57 insertions(+), 56 deletions(-) diff --git a/CvxLean/Command/Solve.lean b/CvxLean/Command/Solve.lean index 84a57a1a..97f1adab 100644 --- a/CvxLean/Command/Solve.lean +++ b/CvxLean/Command/Solve.lean @@ -123,9 +123,10 @@ unsafe def evalSolve : CommandElab := fun stx => addProblemDeclaration (probName ++ `reduced) probReducedOpt false -- Call the solver on prob.reduced and get a point in E. - let coeffsData ← determineCoeffsFromExpr probReducedExpr + let (coeffsData, sections) ← determineCoeffsFromExpr probReducedExpr trace[Meta.debug] "coeffsData: {coeffsData}" - let solPointResponse ← Meta.conicSolverFromValues probReducedExpr coeffsData + let solPointResponse ← + Meta.conicSolverFromValues probReducedExpr coeffsData sections trace[Meta.debug] "solPointResponse: {solPointResponse}" match solPointResponse with diff --git a/CvxLean/Tactic/Solver/Conic.lean b/CvxLean/Tactic/Solver/Conic.lean index b5b7b161..54cb94fa 100644 --- a/CvxLean/Tactic/Solver/Conic.lean +++ b/CvxLean/Tactic/Solver/Conic.lean @@ -26,18 +26,26 @@ def translateCone : ScalarConeType → CBF.ConeType | ScalarConeType.Q => CBF.ConeType.Q | ScalarConeType.QR => CBF.ConeType.QR --- TOOD: Change. This is hacky. -def fixCones : List CBF.Cone → List CBF.Cone - | [] => [] - | List.cons c1 $ List.cons c2 $ List.cons c3 cs => - match c1.type, c2.type, c3.type with - | CBF.ConeType.EXP, CBF.ConeType.EXP, CBF.ConeType.EXP => - (CBF.Cone.mk CBF.ConeType.EXP 3) :: fixCones cs - -- TODO: Generalize to n ≥ 2. - | CBF.ConeType.QR, CBF.ConeType.QR, CBF.ConeType.QR => - (CBF.Cone.mk CBF.ConeType.QR 3) :: fixCones cs - | _, _, _ => c1 :: (fixCones (c2 :: c3 :: cs)) - | List.cons c cs => c :: fixCones cs +def groupCones (sections : ScalarAffineSections) (l : List CBF.Cone) : + MetaM (List CBF.Cone) := do + let l := l.toArray + let mut res := [] + let mut currIdx := 0 + for idx in sections.data do + let group := l[currIdx:idx] + if h : group.size > 0 then + let c := group.get ⟨0, h⟩ + let coneType := c.type + for c' in group do + if !(c'.type = coneType) then + throwError "Only cones of the same type can be grouped." + let totalDim := group.foldl (fun acc c => acc + c.dim) 0 + currIdx := idx + res := res ++ [CBF.Cone.mk coneType totalDim] + else + throwError "Incorrect sections, could not group cones." + + return res /-- -/ def getVars (goalExprs : SolutionExpr) : MetaM (List (Lean.Name × Expr)) := do @@ -56,7 +64,8 @@ unsafe def getTotalDim (goalExprs : SolutionExpr) : MetaM Nat := do return totalDim /-- -/ -unsafe def conicSolverFromValues (goalExprs : SolutionExpr) (data : ProblemData) +unsafe def conicSolverFromValues (goalExprs : SolutionExpr) + (data : ProblemData) (sections : ScalarAffineSections) : MetaM Sol.Response := do let totalDim ← getTotalDim goalExprs @@ -85,12 +94,12 @@ unsafe def conicSolverFromValues (goalExprs : SolutionExpr) (data : ProblemData) let DEnc := CBF.EncodedMatrix.fromArray ma.D cbf := cbf.addMatrixValuedAffineConstraint ma.n HEnc DEnc - -- Fix exponentials. + -- Group cones appropriately, adjusting their dimensions. let n := cbf.scalarConstraints.n let cones := cbf.scalarConstraints.cones - let fixedCones := fixCones cones + let groupedCones ← groupCones sections cones cbf := cbf.setScalarConstraints - (CBF.ConeProduct.mk n fixedCones.length fixedCones) + (CBF.ConeProduct.mk n groupedCones.length groupedCones) -- Write input. let inputPath := "solver/problem.cbf" @@ -132,8 +141,7 @@ unsafe def exprFromSol (goalExprs : SolutionExpr) (sol : Sol.Result) : MetaM Exp -- Generate solution of the correct shape. let solPointExprArrayRaw : Array Expr := - Array.mk $ sol.vars.map (fun v => floatToRealExpr v.activity) - trace[Meta.debug] "raw sol points: {solPointExprArrayRaw}" + Array.mk <| sol.vars.map (fun v => floatToRealExpr v.activity) -- Vectors and matrices as functions. let mut solPointExprArray : Array Expr := #[] @@ -163,7 +171,6 @@ unsafe def exprFromSol (goalExprs : SolutionExpr) (sol : Sol.Result) : MetaM Exp let r ← mkAppM ``List.get! #[arrayList, i''] mkLambdaFVars #[i'] r - trace[Meta.debug] "v: {v}" solPointExprArray := solPointExprArray.push v i := i + n else @@ -188,7 +195,6 @@ unsafe def exprFromSol (goalExprs : SolutionExpr) (sol : Sol.Result) : MetaM Exp let rij ← mkAppM ``List.get! #[ri, j''] mkLambdaFVars #[i', j'] rij - trace[Meta.debug] "M: {M}" solPointExprArray := solPointExprArray.push M i := i + n * m @@ -196,32 +202,6 @@ unsafe def exprFromSol (goalExprs : SolutionExpr) (sol : Sol.Result) : MetaM Exp return solPointExpr --- TODO: Make the tactic work again. -unsafe def conicSolver (goal : MVarId) - : MetaM (List MVarId) := do - let goalExprs ← SolutionExpr.fromGoal goal - let data ← determineCoeffs goal - - let solPointExpr ← conicSolverFromValues goalExprs data - trace[Meta.debug] "conic_solver: {solPointExpr}" - - pure [goal] - end Meta -namespace Tactic - -open Lean.Elab Lean.Elab.Tactic Lean.Meta - -syntax (name := conicSolver) "conic_solver" : tactic - -@[tactic conicSolver] -unsafe def evalConicSolver : Tactic := fun stx => match stx with -| `(tactic| conic_solver) => do - let l ← Meta.conicSolver (← getMainGoal) - replaceMainGoal l -| _ => throwUnsupportedSyntax - -end Tactic - end CvxLean diff --git a/CvxLean/Tactic/Solver/Float/Coeffs.lean b/CvxLean/Tactic/Solver/Float/Coeffs.lean index c801b073..6841f281 100644 --- a/CvxLean/Tactic/Solver/Float/Coeffs.lean +++ b/CvxLean/Tactic/Solver/Float/Coeffs.lean @@ -226,8 +226,13 @@ unsafe def determineMatrixCoeffsAux (e : Expr) (p : Expr) (fty : Expr) instance {n m : ℕ} : OfNat (Fin n.succ ⊕ Fin m.succ) (x) where ofNat := if x <= n then Sum.inl (Fin.ofNat x) else Sum.inr (Fin.ofNat (x - n.succ)) +/-- Indices to group constraints together and tag cones with the correct +dimension when translating the data into CBF format. This happens with the +exponential cone, quadratic cone and rotated quadratic cone, for instance. -/ +def ScalarAffineSections : Type := Array Nat + unsafe def determineCoeffsFromExpr (goalExprs : Meta.SolutionExpr) - : MetaM ProblemData := do + : MetaM (ProblemData × ScalarAffineSections) := do let floatDomain ← realToFloat goalExprs.domain -- Coefficients for objective function. @@ -235,13 +240,16 @@ unsafe def determineCoeffsFromExpr (goalExprs : Meta.SolutionExpr) let objFun ← realToFloat objFun return ← determineScalarCoeffsAux objFun p floatDomain - let constraintsData ← withLambdaBody goalExprs.constraints fun p constraints => do + let (constraintsData, sections) ← + withLambdaBody goalExprs.constraints fun p constraints => do let mut data : ProblemData := ProblemData.empty + let mut sections := #[] -- Constraints without vectors. let cs ← unrollVectors constraints -- Coefficients for constraints. + let mut idx := 0 for c in cs do trace[Meta.debug] "Coeffs going through constraint {c}." match Expr.consumeMData c with @@ -249,10 +257,12 @@ unsafe def determineCoeffsFromExpr (goalExprs : Meta.SolutionExpr) let e ← realToFloat e let res ← determineScalarCoeffsAux e p floatDomain data := data.addZeroConstraint res.1 res.2 + idx := idx + 1 | .app (.const ``Real.posOrthCone _) e => do let e ← realToFloat e let res ← determineScalarCoeffsAux e p floatDomain data := data.addPosOrthConstraint res.1 res.2 + idx := idx + 1 | .app (.app (.app (.const ``Real.expCone _) a) b) c => do let res ← #[a, b, c].mapM fun e => do let e ← realToFloat e @@ -262,6 +272,7 @@ unsafe def determineCoeffsFromExpr (goalExprs : Meta.SolutionExpr) data := data.addExpConstraint res[2]!.1 res[2]!.2 data := data.addExpConstraint res[1]!.1 res[1]!.2 data := data.addExpConstraint res[0]!.1 res[0]!.2 + idx := idx + 3 | .app (.app (.app (.app (.app (.const ``Real.rotatedSoCone _) (.app (.const ``Fin _) n)) _) v) w) x => do let n : Nat ← evalExpr Nat (mkConst ``Nat) n @@ -272,6 +283,7 @@ unsafe def determineCoeffsFromExpr (goalExprs : Meta.SolutionExpr) let e ← realToFloat e let (ea, eb) ← determineScalarCoeffsAux e p floatDomain data := data.addRotatedSOConstraint ea eb + idx := idx + 1 | .app (.app (.app (.app (.const ``Real.soCone _) (.app (.const ``Fin _) n)) _) t) x => do @@ -283,6 +295,7 @@ unsafe def determineCoeffsFromExpr (goalExprs : Meta.SolutionExpr) let e ← realToFloat e let (ea, eb) ← determineScalarCoeffsAux e p floatDomain data := data.addSOConstraint ea eb + idx := idx + 1 -- TODO: Unroll? | .app (.app (.app (.const ``Real.Matrix.posOrthCone _) m) n) e => do let m : Nat ← evalExpr Nat (mkConst ``Nat) m @@ -294,6 +307,7 @@ unsafe def determineCoeffsFromExpr (goalExprs : Meta.SolutionExpr) let eij ← realToFloat eij let res ← determineScalarCoeffsAux eij p floatDomain data := data.addPosOrthConstraint res.1 res.2 + idx := idx + 1 | .app (.app (.app (.const ``Real.Matrix.PSDCone _) mty) _) e => do let e ← realToFloat e @@ -312,17 +326,22 @@ unsafe def determineCoeffsFromExpr (goalExprs : Meta.SolutionExpr) let e ← realToFloat e let (a, b) ← determineScalarCoeffsAux e p floatDomain data := data.addZeroConstraint a b - + idx := idx + 1 | _ => throwError "No match: {c}." - return data + -- New group, add idx. + sections := sections.push idx + return (data, sections) - return constraintsData.setObjectiveOnlyVector objectiveData.1 objectiveData.2 + let (objectiveDataA, objectiveDataB) := objectiveData + let pd := constraintsData.setObjectiveOnlyVector objectiveDataA objectiveDataB + + return (pd, sections) /-- Generate problem data from goal. -/ unsafe def determineCoeffs (goal : Lean.MVarId) : MetaM ProblemData := do let goalExprs ← Meta.SolutionExpr.fromGoal goal - - determineCoeffsFromExpr goalExprs + let (pd, _) ← determineCoeffsFromExpr goalExprs + return pd namespace Tactic diff --git a/CvxLean/Tactic/Solver/Mosek/CBF.lean b/CvxLean/Tactic/Solver/Mosek/CBF.lean index 62506e9d..affeb318 100644 --- a/CvxLean/Tactic/Solver/Mosek/CBF.lean +++ b/CvxLean/Tactic/Solver/Mosek/CBF.lean @@ -25,7 +25,8 @@ end ObjSense /-- Cone type: free, positive orthant, negative orthant, fixpoint zero (for equality), quadratic, quadratic rotated or exponential. -/ inductive ConeType -| F | LPos | LNeg | LEq | Q | QR | EXP + | F | LPos | LNeg | LEq | Q | QR | EXP +deriving DecidableEq namespace ConeType From 8479497c19be55673cd0c12942e40ace80a991fe Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 17:29:21 -0500 Subject: [PATCH 137/189] fix: import in `norm_num` variants --- CvxLean/Tactic/Arith/NormNumVariants.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/CvxLean/Tactic/Arith/NormNumVariants.lean b/CvxLean/Tactic/Arith/NormNumVariants.lean index 3b7b91b4..4913dbec 100644 --- a/CvxLean/Tactic/Arith/NormNumVariants.lean +++ b/CvxLean/Tactic/Arith/NormNumVariants.lean @@ -1,5 +1,6 @@ import Lean import Mathlib.Tactic.NormNum +import CvxLean.Lib.Math.Data.Real namespace CvxLean From 827c296aafcd2f86ef7d9a3b2fe3350ac114e599 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 8 Dec 2023 17:49:11 -0500 Subject: [PATCH 138/189] wip: correct solution for hypersonic shape design --- CvxLean/Test/Convexify/DQCP.lean | 36 +++++++++++++++++++++++++------- 1 file changed, 28 insertions(+), 8 deletions(-) diff --git a/CvxLean/Test/Convexify/DQCP.lean b/CvxLean/Test/Convexify/DQCP.lean index 9c9dc7e0..437c1bdc 100644 --- a/CvxLean/Test/Convexify/DQCP.lean +++ b/CvxLean/Test/Convexify/DQCP.lean @@ -43,23 +43,40 @@ end QCP1 /- -/ section QCP2 +-- def hypersonicShapeDesign (a b : ℝ) := +-- optimization (Δx : ℝ) +-- minimize sqrt (1 / (Δx ^ 2) - 1) +-- subject to +-- h1 : 0 < Δx +-- h2 : Δx ≤ 1 +-- h3 : a * (1 / Δx) - (1 - b) * sqrt (1 - Δx ^ 2) ≤ 0 + +-- time_cmd equivalence redqcp2/dqcp2 : hypersonicShapeDesign 0.05 0.65 := by +-- unfold hypersonicShapeDesign +-- convexify + +-- TODO: the issue comes from the redundant exp constraint that we use +-- as a positivity trick. + def hypersonicShapeDesign (a b : ℝ) := optimization (Δx : ℝ) - minimize sqrt (1 / (Δx ^ 2) - 1) + minimize Δx ^ (-2) - 1 subject to h1 : 0 < Δx h2 : Δx ≤ 1 - h3 : a * (1 / Δx) - (1 - b) * sqrt (1 - Δx ^ 2) ≤ 0 + h3 : (1/7) * (1 / Δx) ≤ sqrt (1 - Δx ^ 2) -time_cmd reduction redqcp2/dqcp2 : hypersonicShapeDesign 0.05 0.65 := by - unfold hypersonicShapeDesign; - convexify - dcp +time_cmd equivalence redqcp2/dqcp2 : hypersonicShapeDesign 0.05 0.65 := by + unfold hypersonicShapeDesign +-- TODO: Error in solve: +-- Lean server printed an error: +-- PANIC at Lean.Expr.mvarId! Lean.Expr:1067:14: mvar expected +-- PANIC at Lean.MetavarContext.getDecl Lean.MetavarContext:398:17: unknown metavariable solve dqcp2 -#print dqcp2 --- def dqcp2 : Minimization (ℝ × ℝ × ℝ × ℝ × ℝ × ℝ × ℝ × ℝ) ℝ := +#print dqcp2.reduced +-- def dqcp2.reduced : Minimization (ℝ × ℝ × ℝ × ℝ × ℝ × ℝ × ℝ × ℝ) ℝ := -- optimization (Δx : ℝ) (t₀.0 : ℝ) (t₁.1 : ℝ) (t.2 : ℝ) (y'.3 : ℝ) (t.0.4 : ℝ) (t.5 : ℝ) (t.6 : ℝ) -- minimize t₀.0 - 1 -- subject to @@ -77,6 +94,9 @@ solve dqcp2 -- _ : rotatedSoCone t.5 (1 / 2) ![Δx] -- _ : rotatedSoCone (1 - t.5) (1 / 2) ![t.6] + + -- _ : posOrthCone (t.6 - t.2 / 20 / (7 / 20)) + -- TODO: this is wrong. #eval dqcp2.value #eval dqcp2.solution From a3edccf8ec783e4a26155d8c114d71fbb498a7bd Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Sat, 9 Dec 2023 14:33:30 -0500 Subject: [PATCH 139/189] chore: note on solver stalling --- CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean | 3 ++- CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean | 1 + 2 files changed, 3 insertions(+), 1 deletion(-) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean index 7b292fd9..31da8d07 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean @@ -18,7 +18,8 @@ implementationVars (t : ℝ) (y' : ℝ) implementationObjective (y - x - t) implementationConstraints (c1 : expCone t x y) - -- NOTE: This is a trick to make y strictly positive. + -- NOTE(RFM): This is a trick to make y strictly positive. + -- TODO(RFM): This can make the solver stall (error 166). (c2 : exp y' ≤ y) solution (t := -(x * log (x / y))) (y' := log y) solutionEqualsAtom by diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean index 3c8e33cb..07628754 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean @@ -23,6 +23,7 @@ implementationConstraints (c1 : soCone (y + t) ![y - t, 2 * x]) (c2 : 0 ≤ t) -- NOTE(RFM): This is a trick to make y strictly positive. + -- TODO(RFM): This makes the solver stall (error 166). (c3 : exp y' ≤ y) solution (t := x ^ 2 / y) (y' := log y) solutionEqualsAtom by rfl From eadf336b15541f441e88d1681da5531b981fbaed Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Sat, 9 Dec 2023 15:15:15 -0500 Subject: [PATCH 140/189] test: more convexify Rust tests --- egg-convexify/tests/test_dqcp.rs | 14 +++++++++++++- egg-convexify/tests/test_misc.rs | 13 +++++++++++++ 2 files changed, 26 insertions(+), 1 deletion(-) diff --git a/egg-convexify/tests/test_dqcp.rs b/egg-convexify/tests/test_dqcp.rs index 38f6894e..3aa3485e 100644 --- a/egg-convexify/tests/test_dqcp.rs +++ b/egg-convexify/tests/test_dqcp.rs @@ -3,12 +3,13 @@ Tests from quasiconvex programming. !*/ use egg_convexify::domain; +use domain::Domain as Domain; use egg_convexify::test_util::{*}; #[test] -fn test_qcp() { +fn test_qcp1() { convexify_check_with_domain( vec![("x", domain::pos_dom())], "(var x)", @@ -16,3 +17,14 @@ fn test_qcp() { "(le (sqrt (div (var x) (add (var x) 1))) 1)" ]); } + +#[test] +fn test_qcp2() { + let d = Domain::make_oc(domain::zero(), domain::one()); + convexify_check_with_domain( + vec![("x", d)], + "(sqrt (sub (div 1 (pow (var x) 2)) 1))", + vec![ + "(le (sub (mul (div 1 20) (div 1 (var x))) (mul (div 7 20) (sqrt (sub 1 (pow (var x) 2))))) 0)" + ]) +} diff --git a/egg-convexify/tests/test_misc.rs b/egg-convexify/tests/test_misc.rs index e9c2cc5d..71ee2ccc 100644 --- a/egg-convexify/tests/test_misc.rs +++ b/egg-convexify/tests/test_misc.rs @@ -19,3 +19,16 @@ fn test_sqrt_pow4() { vec![("x", domain::nonneg_dom())], "(sqrt (pow (var x) 4))"); } + +#[test] +fn test_div_constant_simp() { + convexify_check_expression( + "(div (div (var x) 20) (div 7 20))"); +} + +#[test] +fn test_div_constant_le_simp() { + convexify_check_expression_with_domain( + vec![("x", domain::nonneg_dom()), ("y", domain::pos_dom())], + "(le (div (qol 1 (var y)) 20) (mul (div 7 20) (sqrt (var x))))"); +} From be94ac60fd3431f45e3f2c91b08c58a89f374b05 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Sat, 9 Dec 2023 15:22:36 -0500 Subject: [PATCH 141/189] feat: more division rules --- .../Tactic/Convexify/RewriteMapLibrary.lean | 22 +++++++++++++++++-- egg-convexify/src/rules.rs | 17 ++++++++++++++ 2 files changed, 37 insertions(+), 2 deletions(-) diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index 37076ba7..ee5b775e 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -132,6 +132,9 @@ register_rewrite_map "mul_sub" ; "(mul ?a (sub ?b ?c))" => "(sub (mul ?a ?b) (mu register_rewrite_map "mul_sub-rev" ; "(sub (mul ?a ?b) (mul ?a ?c))" => "(mul ?a (sub ?b ?c))" := simp_or_rw [←mul_sub]; +register_rewrite_map "div_one" ; "(div ?a 1)" => "?a" := + simp_or_rw [div_one]; + register_rewrite_map "add_div" ; "(div (add ?a ?b) ?c)" => "(add (div ?a ?c) (div ?b ?c))" := simp_or_rw [add_div]; @@ -144,6 +147,21 @@ register_rewrite_map "mul_div" ; "(mul ?a (div ?b ?c))" => "(div (mul ?a ?b) ?c) register_rewrite_map "mul_div-rev" ; "(div (mul ?a ?b) ?c)" => "(mul ?a (div ?b ?c))" := simp_or_rw [←mul_div (G := ℝ)]; +register_rewrite_map "div_mul" ; "(mul (div ?a ?b) ?c)" => "(div ?a (div ?b ?c))" := + simp_or_rw [div_mul]; + +register_rewrite_map "div_mul-rev" ; "(div ?a (div ?b ?c))" => "(mul (div ?a ?b) ?c)" := + simp_or_rw [←div_mul]; + +register_rewrite_map "mul_div_mul_left" ; "(div (mul ?c ?a) (mul ?c ?b))" => "(div ?a ?b)" := + simp_or_rw [mul_div_mul_left (G₀ := ℝ) _ _ (by positivity!)]; + +register_rewrite_map "div_div" ; "(div (div ?a ?b) ?c)" => "(div ?a (mul ?b ?c))" := + simp_or_rw [div_div]; + +register_rewrite_map "div_div-rev" ; "(div ?a (mul ?b ?c))" => "(div (div ?a ?b) ?c)" := + simp_or_rw [←div_div]; + register_rewrite_map "div_self" ; "(div ?a ?a)" => "1" := simp_or_rw [div_self (G₀ := ℝ) (by positivity!)]; @@ -216,7 +234,7 @@ register_rewrite_map "pow_two-rev"; "(mul ?a ?a)" => "(pow ?a 2)" := register_rewrite_map "pow_half_two"; "(pow (pow ?a 0.5) 2)" => "?a" := simp_or_rw [Real.pow_half_two (by positivity!)]; --- TODO(RFM): Technically ← but no pattern to match on otherwise. +-- Exception, technically ← but we cannot find the pattern otherwise. register_rewrite_map "pow_half_two-rev"; "?a" => "(pow (pow ?a 0.5) 2)" := simp_or_rw [Real.pow_half_two (by positivity!)]; @@ -271,7 +289,7 @@ register_rewrite_map "log_pow" ; "(log (pow ?a ?b))" => "(mul ?b (log ?a))" := register_rewrite_map "log_pow-rev" ; "(mul ?b (log ?a))" => "(log (pow ?a ?b))" := simp_or_rw [←Real.log_rpow (by positivity!)]; --- NOTE: Special rule that only works if the exponent is a natural number. +-- Special type of rule as it only works if the exponent is a natural number. register_rewrite_map "log_pow_nat" ; "(log (pow ?a ?b))" => "(mul ?b (log ?a))" := (norm_cast; simp_or_rw [Real.log_pow]); diff --git a/egg-convexify/src/rules.rs b/egg-convexify/src/rules.rs index f4655ec1..b8c39076 100644 --- a/egg-convexify/src/rules.rs +++ b/egg-convexify/src/rules.rs @@ -100,6 +100,8 @@ pub fn rules() -> Vec> { vec![ rw!("mul_sub-rev"; "(sub (mul ?a ?b) (mul ?a ?c))" => "(mul ?a (sub ?b ?c))"), + rw!("div_one"; "(div ?a 1)" => "?a"), + rw!("add_div"; "(div (add ?a ?b) ?c)" => "(add (div ?a ?c) (div ?b ?c))"), rw!("add_div-rev"; "(add (div ?a ?c) (div ?b ?c))" => "(div (add ?a ?b) ?c)"), @@ -108,6 +110,21 @@ pub fn rules() -> Vec> { vec![ rw!("mul_div-rev"; "(div (mul ?a ?b) ?c)" => "(mul ?a (div ?b ?c))"), + rw!("div_mul"; "(mul (div ?a ?b) ?c)" => "(div ?a (div ?b ?c))" + if is_not_zero("?b") if is_not_zero("?c")), + + rw!("div_mul-rev"; "(div ?a (div ?b ?c))" => "(mul (div ?a ?b) ?c)" + if is_not_zero("?b") if is_not_zero("?c")), + + rw!("mul_div_mul_left"; "(div (mul ?c ?a) (mul ?c ?b))" => "(div ?a ?b)" + if is_not_zero("?c")), + + rw!("div_div"; "(div (div ?a ?b) ?c)" => "(div ?a (mul ?b ?c))" + if is_not_zero("?b") if is_not_zero("?c")), + + rw!("div_div-rev"; "(div ?a (mul ?b ?c))" => "(div (div ?a ?b) ?c)" + if is_not_zero("?b") if is_not_zero("?c")), + rw!("div_self"; "(div ?a ?a)" => "1" if is_not_zero("?a")), rw!("inv_eq_one_div"; "(inv ?a)" => "(div 1 ?a)" if is_not_zero("?a")), From f044dc1d4937359183db9c7265f8957d933e22a0 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Sat, 9 Dec 2023 16:16:52 -0500 Subject: [PATCH 142/189] feat: `sqNat` --- CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean | 21 +++++++++++++++++++++ 1 file changed, 21 insertions(+) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean index 09cf5094..5e7073ad 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean @@ -28,6 +28,27 @@ optimality by exact h vconditionElimination +declare_atom sqNat [convex] (x : ℝ)? : x ^ (2 : ℕ) := +vconditions +implementationVars (t : ℝ) +implementationObjective (t) +implementationConstraints + (c1 : rotatedSoCone t (1/2) ![x]) +solution + (t := x ^ 2) +solutionEqualsAtom rfl +feasibility + (c1 : by + unfold rotatedSoCone + simp + exact ⟨sq_nonneg x, zero_le_two⟩) +optimality by + unfold rotatedSoCone at c1 + have h := c1.1 + simp at h ⊢ + exact h +vconditionElimination + declare_atom Vec.sq [convex] (n : ℕ)& (x : Fin n → ℝ)? : x ^ 2 := vconditions implementationVars (t : Fin n → ℝ) From c00913a1b0dd64c2f5fa64caf518735b267eee83 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Sat, 9 Dec 2023 16:53:59 -0500 Subject: [PATCH 143/189] fix: constraints with the same name in convexify --- CvxLean/Tactic/Convexify/Convexify.lean | 8 +- .../Convexify/Egg/FromMinimization.lean | 82 +++++++++++-------- 2 files changed, 52 insertions(+), 38 deletions(-) diff --git a/CvxLean/Tactic/Convexify/Convexify.lean b/CvxLean/Tactic/Convexify/Convexify.lean index c39e16c5..9f052ab7 100644 --- a/CvxLean/Tactic/Convexify/Convexify.lean +++ b/CvxLean/Tactic/Convexify/Convexify.lean @@ -96,7 +96,11 @@ def evalStep (g : MVarId) (step : EggRewrite) (vars : List Name) (fvars : Array Expr) (domain : Expr) (numConstrTags : ℕ) (tagsMap : HashMap String ℕ) (isEquiv : Bool) : TacticM (List MVarId) := withMainContext <| do - let tag := step.location + let tag ← liftMetaM <| do + if let [_, tag] := step.location.splitOn ":" then + return tag + else + throwError "Unexpected tag name {step.location}." let tagNum := tagsMap.find! tag let atObjFun := tagNum == 0 @@ -261,7 +265,7 @@ def evalConvexify : Tactic := fun stx => match stx with let varDomainConstrs := domainConstrs.map (fun (_, v, d) => (v, d)) let constrsToIgnore := domainConstrs.map (fun (h, _, _) => h) - -- Remove less-than's before sending it to egg. + -- Remove domain constraints before sending it to egg. let gStr := { gStr with constr := gStr.constr.filter (fun (h, _) => !constrsToIgnore.contains h) } diff --git a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean index cc0a6bb9..631eda98 100644 --- a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean +++ b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean @@ -158,42 +158,50 @@ def _root_.ExtendedEggTree.fromMinimization (e : Meta.MinimizationExpr) (vars : -- * x < f and f < x, -- * x = f and f = x, -- where "x" is a variable name and "f" is a numerical value. - let domainConstrs := ocTree.constr.filterMap <| fun (h, c) => - match c with - | Tree.node "le" #[(tl : EggTree), (tr : EggTree)] => - match (tl.getVariableName? vars, tr.getNumericalValue?) with - | (some v, some f) => - -- v ∈ (-inf, f] - some (h, v, ⟨"-inf", f.toString, "1", "0"⟩) - | _ => - match (tl.getNumericalValue?, tr.getVariableName? vars) with - | (some f, some v) => - -- v ∈ [f, inf) - some (h, v, ⟨f.toString, "inf", "0", "1"⟩) - | _ => none - | Tree.node "lt" #[(tl : EggTree), (tr : EggTree)] => - match (tl.getVariableName? vars, tr.getNumericalValue?) with - | (some v, some f) => - -- v ∈ (-inf, f) - some (h, v, ⟨"-inf", f.toString, "1", "1"⟩) - | _ => - match (tl.getNumericalValue?, tr.getVariableName? vars) with - | (some f, some v) => - -- v ∈ (f, inf) - some (h, v, ⟨f.toString, "inf", "1", "1"⟩) - | _ => none - | Tree.node "eq" #[(tl : EggTree), (tr : EggTree)] => - match (tl.getVariableName? vars, tr.getNumericalValue?) with - | (some v, some f) => - -- v ∈ [f, f] - some (h, v, ⟨f.toString, f.toString, "1", "1"⟩) - | _ => - match (tl.getNumericalValue?, tr.getVariableName? vars) with - | (some f, some v) => + let mut domainConstrs := #[] + let mut i := 0 + for (h, c) in ocTree.constr do + -- TODO: some constraints may have the same name, so we add the index. + let h := s!"{i}:" ++ h + let res := + match c with + | Tree.node "le" #[(tl : EggTree), (tr : EggTree)] => + match (tl.getVariableName? vars, tr.getNumericalValue?) with + | (some v, some f) => + -- v ∈ (-inf, f] + some (h, v, ⟨"-inf", f.toString, "1", "0"⟩) + | _ => + match (tl.getNumericalValue?, tr.getVariableName? vars) with + | (some f, some v) => + -- v ∈ [f, inf) + some (h, v, ⟨f.toString, "inf", "0", "1"⟩) + | _ => none + | Tree.node "lt" #[(tl : EggTree), (tr : EggTree)] => + match (tl.getVariableName? vars, tr.getNumericalValue?) with + | (some v, some f) => + -- v ∈ (-inf, f) + some (h, v, ⟨"-inf", f.toString, "1", "1"⟩) + | _ => + match (tl.getNumericalValue?, tr.getVariableName? vars) with + | (some f, some v) => + -- v ∈ (f, inf) + some (h, v, ⟨f.toString, "inf", "1", "1"⟩) + | _ => none + | Tree.node "eq" #[(tl : EggTree), (tr : EggTree)] => + match (tl.getVariableName? vars, tr.getNumericalValue?) with + | (some v, some f) => -- v ∈ [f, f] - some (h, v, ⟨f.toString, f.toString, "1", "1"⟩) - | _ => none - | _ => none + some (h, v, ⟨f.toString, f.toString, "0", "0"⟩) + | _ => + match (tl.getNumericalValue?, tr.getVariableName? vars) with + | (some f, some v) => + -- v ∈ [f, f] + some (h, v, ⟨f.toString, f.toString, "0", "0"⟩) + | _ => none + | _ => none + match res with + | some r => domainConstrs := domainConstrs.push r + | none => pure () -- Surround variables and adjust operations in the whole tree. let ocTree : OC (String × Tree String String) := { @@ -201,7 +209,9 @@ def _root_.ExtendedEggTree.fromMinimization (e : Meta.MinimizationExpr) (vars : constr := ocTree.constr.map (fun (h, c) => (h, EggTree.surroundVars c vars)) } let ocTree : OC (String × Tree String String) := { objFun := (ocTree.objFun.1, ← EggTree.adjustOps ocTree.objFun.2) - constr := ← ocTree.constr.mapM (fun (h, c) => return (h, ← EggTree.adjustOps c)) } + constr := ← ocTree.constr.mapIdxM <| + -- TODO: some constraints may have the same name, so we add the index. + fun i (h, c) => return (s!"{i}:" ++ h, ← EggTree.adjustOps c) } return (ocTree, domainConstrs) end Egg.FromMinimization From d21b6224a168b9e968ff28a233b1326108ddb99b Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Sat, 9 Dec 2023 16:59:47 -0500 Subject: [PATCH 144/189] feat: correct solution for HSD but sorrys in QOL --- CvxLean/Tactic/Convexify/Convexify.lean | 6 ++- .../DCP/AtomLibrary/Fns/QuadOverLin.lean | 15 +++---- CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean | 21 ---------- CvxLean/Test/Convexify/DQCP.lean | 39 +++++-------------- 4 files changed, 22 insertions(+), 59 deletions(-) diff --git a/CvxLean/Tactic/Convexify/Convexify.lean b/CvxLean/Tactic/Convexify/Convexify.lean index 9f052ab7..fefefbf1 100644 --- a/CvxLean/Tactic/Convexify/Convexify.lean +++ b/CvxLean/Tactic/Convexify/Convexify.lean @@ -97,11 +97,13 @@ def evalStep (g : MVarId) (step : EggRewrite) (numConstrTags : ℕ) (tagsMap : HashMap String ℕ) (isEquiv : Bool) : TacticM (List MVarId) := withMainContext <| do let tag ← liftMetaM <| do - if let [_, tag] := step.location.splitOn ":" then + if step.location == "objFun" then + return "objFun" + else if let [_, tag] := step.location.splitOn ":" then return tag else throwError "Unexpected tag name {step.location}." - let tagNum := tagsMap.find! tag + let tagNum := tagsMap.find! step.location let atObjFun := tagNum == 0 -- Special case when mapping the objective function in equivalence mode. Note diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean index 07628754..4950d4a2 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean @@ -24,7 +24,7 @@ implementationConstraints (c2 : 0 ≤ t) -- NOTE(RFM): This is a trick to make y strictly positive. -- TODO(RFM): This makes the solver stall (error 166). - (c3 : exp y' ≤ y) + -- (c3 : exp y' ≤ y) solution (t := x ^ 2 / y) (y' := log y) solutionEqualsAtom by rfl feasibility @@ -38,11 +38,11 @@ feasibility have hynn := le_of_lt hy rw [rpow_two] exact div_nonneg (pow_two_nonneg x) hynn) - (c3 : by - simp [exp_log hy]) + -- (c3 : by + -- simp [exp_log hy]) optimality by intros z hyz - have hy := lt_of_lt_of_le (exp_pos _) c3 + have hy : 0 < y := sorry --lt_of_lt_of_le (exp_pos _) c3 have hynn := le_of_lt hy rw [soCone_add_sub_two_mul_of_nonneg x hynn c2] at c1 have hz := lt_of_lt_of_le hy hyz @@ -51,8 +51,9 @@ optimality by apply mul_le_mul hyz (le_refl _) c2 (le_of_lt hz) vconditionElimination (hy : by - intros z hz - have hy := lt_of_lt_of_le (exp_pos _) c3 - exact lt_of_lt_of_le hy hz) + -- intros z hz + -- have hy := lt_of_lt_of_le (exp_pos _) c3 + -- exact lt_of_lt_of_le hy hz + sorry) end CvxLean diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean index 5e7073ad..09cf5094 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean @@ -28,27 +28,6 @@ optimality by exact h vconditionElimination -declare_atom sqNat [convex] (x : ℝ)? : x ^ (2 : ℕ) := -vconditions -implementationVars (t : ℝ) -implementationObjective (t) -implementationConstraints - (c1 : rotatedSoCone t (1/2) ![x]) -solution - (t := x ^ 2) -solutionEqualsAtom rfl -feasibility - (c1 : by - unfold rotatedSoCone - simp - exact ⟨sq_nonneg x, zero_le_two⟩) -optimality by - unfold rotatedSoCone at c1 - have h := c1.1 - simp at h ⊢ - exact h -vconditionElimination - declare_atom Vec.sq [convex] (n : ℕ)& (x : Fin n → ℝ)? : x ^ 2 := vconditions implementationVars (t : Fin n → ℝ) diff --git a/CvxLean/Test/Convexify/DQCP.lean b/CvxLean/Test/Convexify/DQCP.lean index 437c1bdc..5aaf3628 100644 --- a/CvxLean/Test/Convexify/DQCP.lean +++ b/CvxLean/Test/Convexify/DQCP.lean @@ -43,45 +43,32 @@ end QCP1 /- -/ section QCP2 --- def hypersonicShapeDesign (a b : ℝ) := --- optimization (Δx : ℝ) --- minimize sqrt (1 / (Δx ^ 2) - 1) --- subject to --- h1 : 0 < Δx --- h2 : Δx ≤ 1 --- h3 : a * (1 / Δx) - (1 - b) * sqrt (1 - Δx ^ 2) ≤ 0 - --- time_cmd equivalence redqcp2/dqcp2 : hypersonicShapeDesign 0.05 0.65 := by --- unfold hypersonicShapeDesign --- convexify - --- TODO: the issue comes from the redundant exp constraint that we use --- as a positivity trick. - def hypersonicShapeDesign (a b : ℝ) := optimization (Δx : ℝ) - minimize Δx ^ (-2) - 1 + minimize sqrt (1 / (Δx ^ 2) - 1) subject to h1 : 0 < Δx h2 : Δx ≤ 1 - h3 : (1/7) * (1 / Δx) ≤ sqrt (1 - Δx ^ 2) + h3 : a * (1 / Δx) - (1 - b) * sqrt (1 - Δx ^ 2) ≤ 0 +set_option trace.Meta.debug true in time_cmd equivalence redqcp2/dqcp2 : hypersonicShapeDesign 0.05 0.65 := by unfold hypersonicShapeDesign + convexify --- TODO: Error in solve: +-- TODO: Error in solve (or dcp): -- Lean server printed an error: -- PANIC at Lean.Expr.mvarId! Lean.Expr:1067:14: mvar expected -- PANIC at Lean.MetavarContext.getDecl Lean.MetavarContext:398:17: unknown metavariable solve dqcp2 #print dqcp2.reduced --- def dqcp2.reduced : Minimization (ℝ × ℝ × ℝ × ℝ × ℝ × ℝ × ℝ × ℝ) ℝ := --- optimization (Δx : ℝ) (t₀.0 : ℝ) (t₁.1 : ℝ) (t.2 : ℝ) (y'.3 : ℝ) (t.0.4 : ℝ) (t.5 : ℝ) (t.6 : ℝ) +-- def dqcp2.reduced : Minimization (ℝ × ℝ × ℝ × ℝ × ℝ × ℝ × ℝ) ℝ := +-- optimization (Δx : ℝ) (t₀.0 : ℝ) (t₁.1 : ℝ) (t.2 : ℝ) (y'.3 : ℝ) (t.4 : ℝ) (t.5 : ℝ) -- minimize t₀.0 - 1 -- subject to -- _ : posOrthCone (1 - Δx) --- _ : posOrthCone (0 - 7 / 20 * (t.2 - t.6)) +-- _ : posOrthCone (7 / 20 * t.5 - t.2 / 20) -- _ : soCone (Δx + t₁.1) ![Δx - t₁.1, 2] -- _ : soCone (t₀.0 + 1) ![t₀.0 - 1, 2 * t₁.1] -- _ : posOrthCone (t₀.0 - 0) @@ -89,15 +76,9 @@ solve dqcp2 -- _ : posOrthCone (Δx - 0) -- _ : soCone (Δx + t.2) ![Δx - t.2, 2 * 1] -- _ : posOrthCone (t.2 - 0) --- _ : posOrthCone (Δx - t.0.4) --- _ : expCone y'.3 1 t.0.4 --- _ : rotatedSoCone t.5 (1 / 2) ![Δx] --- _ : rotatedSoCone (1 - t.5) (1 / 2) ![t.6] - - - -- _ : posOrthCone (t.6 - t.2 / 20 / (7 / 20)) +-- _ : rotatedSoCone t.4 (1 / 2) ![Δx] +-- _ : rotatedSoCone (1 - t.4) (1 / 2) ![t.5] --- TODO: this is wrong. #eval dqcp2.value #eval dqcp2.solution From 1be6d20b954c57c92fb7b6ca01ab33775ba2603f Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Sat, 9 Dec 2023 17:09:41 -0500 Subject: [PATCH 145/189] =?UTF-8?q?fix:=20remove=20`exp=20y'=20=E2=89=A4?= =?UTF-8?q?=20y`=20from=20everywhere?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean | 82 ++++++++----------- .../DCP/AtomLibrary/Fns/QuadOverLin.lean | 26 +++--- 2 files changed, 48 insertions(+), 60 deletions(-) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean index 31da8d07..307d6f38 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean @@ -10,18 +10,18 @@ namespace CvxLean open Real +-- TODO(RFM): same issue as in `quadOverLin`, no clear way to have 0 < y as a +-- vcondition, the exp y' ≤ y trick makes the solver stall. declare_atom klDiv [convex] (x : ℝ)? (y : ℝ)? : x * log (x / y) - x + y := vconditions (hx : 0 ≤ x) - (hy : 0 < y) -implementationVars (t : ℝ) (y' : ℝ) + (hy : 10e-6 ≤ y) +implementationVars (t : ℝ) implementationObjective (y - x - t) implementationConstraints (c1 : expCone t x y) - -- NOTE(RFM): This is a trick to make y strictly positive. - -- TODO(RFM): This can make the solver stall (error 166). - (c2 : exp y' ≤ y) -solution (t := -(x * log (x / y))) (y' := log y) + (c2 : 10e-6 ≤ y) +solution (t := -(x * log (x / y))) solutionEqualsAtom by ring feasibility @@ -32,15 +32,17 @@ feasibility | inl hx => left refine ⟨hx, ?_⟩ + have hypos : 0 < y := by positivity rw [mul_comm _ (log _), ←neg_mul, ←mul_div, div_self (ne_of_gt hx)] - rw [mul_one, exp_neg, exp_log (div_pos hx hy), inv_div, mul_div] + rw [mul_one, exp_neg, exp_log (div_pos hx hypos), inv_div, mul_div] rw [div_le_iff hx, mul_comm] | inr hx => replace hx := hx.symm right - simp [hx, le_of_lt hy]) + simp [hx] + positivity) (c2 : by - simp [exp_log hy]) + exact hy) optimality by unfold expCone at c1 cases c1 with @@ -66,18 +68,18 @@ vconditionElimination | inl c => exact le_of_lt c.1 | inr c => rw [c.1]) (hy : by - exact lt_of_lt_of_le (exp_pos _) c2) + exact c2) declare_atom klDiv2 [convex] (x : ℝ)? (y : ℝ)? : klDiv x y := vconditions (hx : 0 ≤ x) - (hy : 0 < y) -implementationVars (t : ℝ) (y' : ℝ) + (hy : 10e-6 ≤ y) +implementationVars (t : ℝ) implementationObjective (y - x - t) implementationConstraints (c1 : expCone t x y) - (c2 : exp y' ≤ y) -solution (t := -(x * log (x / y))) (y' := log y) + (c2 : 10e-6 ≤ y) +solution (t := -(x * log (x / y))) solutionEqualsAtom by unfold klDiv ring @@ -89,36 +91,30 @@ feasibility unfold posOrthCone at h simpa using h) optimality by - apply klDiv.optimality x y t y' (exp y') c1 + apply klDiv.optimality x y t c1 { unfold posOrthCone expCone at * simpa using c2 } - { unfold expCone at * - simp } vconditionElimination (hx : by - apply klDiv.vcondElim0 x y t y' (exp y') c1 - { unfold posOrthCone expCone at * - simpa using c2 } - { unfold expCone at * - simp }) + apply klDiv.vcondElim0 x y t c1 + unfold posOrthCone at * + simpa using c2) (hy : by - apply klDiv.vcondElim1 x y t y' (exp y') c1 - { unfold posOrthCone expCone at * - simpa using c2 } - { unfold expCone at * - simp }) + apply klDiv.vcondElim1 x y t c1 + unfold posOrthCone expCone at * + simpa using c2) declare_atom Vec.klDiv [convex] (m : Nat)& (x : Fin m → ℝ)? (y : Fin m → ℝ)? : Vec.klDiv x y := vconditions (hx : 0 ≤ x) - (hy : ∀ i, 0 < y i) -implementationVars (t : Fin m → ℝ) (y' : Fin m → ℝ) + (hy : (fun _ => 10e-6) ≤ y) +implementationVars (t : Fin m → ℝ) implementationObjective (y - x - t) implementationConstraints (c1 : Vec.expCone t x y) - (c2 : Vec.exp y' ≤ y) -solution (t := fun i => -((x i) * log ((x i) / (y i)))) (y' := Vec.log y) + (c2 : Vec.const m 10e-6 ≤ y) +solution (t := fun i => -((x i) * log ((x i) / (y i)))) solutionEqualsAtom by simp [Vec.klDiv, klDiv]; ext i; simp; ring feasibility @@ -135,25 +131,19 @@ feasibility simpa using h) optimality fun i => by unfold Vec.expCone at c1 - apply klDiv.optimality (x i) (y i) (t i) (y' i) (exp (y' i)) (c1 i) - { unfold posOrthCone expCone at * - simpa using (c2 i) } - { unfold expCone at * - simp } + apply klDiv.optimality (x i) (y i) (t i) (c1 i) + unfold posOrthCone expCone at * + simpa using (c2 i) vconditionElimination (hx : fun i => by unfold Vec.expCone at c1 - apply klDiv.vcondElim0 (x i) (y i) (t i) (y' i) (exp (y' i)) (c1 i) - { unfold posOrthCone expCone at * - simpa using (c2 i) } - { unfold expCone at * - simp }) + apply klDiv.vcondElim0 (x i) (y i) (t i) (c1 i) + unfold posOrthCone expCone at * + simpa using (c2 i)) (hy : fun i => by unfold Vec.expCone at c1 - apply klDiv.vcondElim1 (x i) (y i) (t i) (y' i) (exp (y' i)) (c1 i) - { unfold posOrthCone expCone at * - simpa using (c2 i) } - { unfold expCone at * - simp }) + apply klDiv.vcondElim1 (x i) (y i) (t i) (c1 i) + unfold posOrthCone expCone at * + simpa using (c2 i)) end CvxLean diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean index 4950d4a2..58b05b6f 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean @@ -14,35 +14,35 @@ open Real -- TODO(RFM): generalize to x : Fin n → ℝ -- TODO(RFM): distinguish between nonincreasing and nondecreasing. +-- TODO(RFM): how to do strict positivity and make vcondelimination work? declare_atom quadOverLin [convex] (x : ℝ)? (y : ℝ)- : x ^ 2 / y := vconditions - (hy : 0 < y) + (hy : 10e-6 ≤ y) implementationVars (t : ℝ) (y' : ℝ) implementationObjective (t) implementationConstraints (c1 : soCone (y + t) ![y - t, 2 * x]) (c2 : 0 ≤ t) - -- NOTE(RFM): This is a trick to make y strictly positive. - -- TODO(RFM): This makes the solver stall (error 166). - -- (c3 : exp y' ≤ y) + (c3 : 10e-6 ≤ y) solution (t := x ^ 2 / y) (y' := log y) solutionEqualsAtom by rfl feasibility (c1 : by - have hynn := le_of_lt hy + have hypos : 0 < y := by positivity + have hynn : 0 ≤ y := by linarith have hx2ynn : 0 ≤ x ^ 2 / y := by rw [rpow_two]; exact div_nonneg (pow_two_nonneg x) hynn rw [soCone_add_sub_two_mul_of_nonneg _ hynn hx2ynn] - rw [mul_div, mul_comm, ←mul_div, div_self (ne_of_gt hy), mul_one]) + rw [mul_div, mul_comm, ←mul_div, div_self (ne_of_gt hypos), mul_one]) (c2 : by - have hynn := le_of_lt hy + have hynn : 0 ≤ y := by positivity rw [rpow_two] exact div_nonneg (pow_two_nonneg x) hynn) - -- (c3 : by - -- simp [exp_log hy]) + (c3 : by + exact hy) optimality by intros z hyz - have hy : 0 < y := sorry --lt_of_lt_of_le (exp_pos _) c3 + have hy : 0 < y := by positivity have hynn := le_of_lt hy rw [soCone_add_sub_two_mul_of_nonneg x hynn c2] at c1 have hz := lt_of_lt_of_le hy hyz @@ -51,9 +51,7 @@ optimality by apply mul_le_mul hyz (le_refl _) c2 (le_of_lt hz) vconditionElimination (hy : by - -- intros z hz - -- have hy := lt_of_lt_of_le (exp_pos _) c3 - -- exact lt_of_lt_of_le hy hz - sorry) + intros z hz + linarith) end CvxLean From 49a32bd06e871eb109e59d1221f3cbe59242f0cd Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 10:58:33 -0500 Subject: [PATCH 146/189] fix: increment index in `fromMinimization` --- CvxLean/Tactic/Convexify/Egg/FromMinimization.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean index 631eda98..5b7de02a 100644 --- a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean +++ b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean @@ -161,7 +161,7 @@ def _root_.ExtendedEggTree.fromMinimization (e : Meta.MinimizationExpr) (vars : let mut domainConstrs := #[] let mut i := 0 for (h, c) in ocTree.constr do - -- TODO: some constraints may have the same name, so we add the index. + -- NOTE: some constraints may have the same name, so we add the index. let h := s!"{i}:" ++ h let res := match c with @@ -202,6 +202,7 @@ def _root_.ExtendedEggTree.fromMinimization (e : Meta.MinimizationExpr) (vars : match res with | some r => domainConstrs := domainConstrs.push r | none => pure () + i := i + 1 -- Surround variables and adjust operations in the whole tree. let ocTree : OC (String × Tree String String) := { From 09bf45766a6088d131b5d26c0ffddf96f582c139 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 11:02:48 -0500 Subject: [PATCH 147/189] fix: do not print "Empty domain" --- egg-convexify/src/optimization.rs | 3 +- egg-convexify/src/rules.rs | 92 +++++++++++++++---------------- 2 files changed, 47 insertions(+), 48 deletions(-) diff --git a/egg-convexify/src/optimization.rs b/egg-convexify/src/optimization.rs index dd60bbe6..99e4da74 100644 --- a/egg-convexify/src/optimization.rs +++ b/egg-convexify/src/optimization.rs @@ -65,8 +65,7 @@ impl Analysis for Meta { if !d_to.eq(&d_from) { let inter = d_to.intersection(&d_from); if Domain::is_empty(&inter) { - // Should never get here. - println!("Empty domain."); + // Should never get here for feasible problems. to.domain = None } else { to.domain = Some(inter); diff --git a/egg-convexify/src/rules.rs b/egg-convexify/src/rules.rs index b8c39076..379fb222 100644 --- a/egg-convexify/src/rules.rs +++ b/egg-convexify/src/rules.rs @@ -13,7 +13,7 @@ pub fn rules() -> Vec> { vec![ /* Objective function rules. */ - rw!("map_objFun_log"; "(objFun ?a)" => "(objFun (log ?a))" + rw!("map_objFun_log"; "(objFun ?a)" => "(objFun (log ?a))" if is_gt_zero("?a")), rw!("map_objFun_sq"; "(objFun ?a)" => "(objFun (pow ?a 2))" @@ -21,40 +21,40 @@ pub fn rules() -> Vec> { vec![ /* Equality rules. */ - // NOTE: many more rules could apply here, but in our examples, equalities + // NOTE: many more rules could apply here, but in our examples, equalities // were either already affine or required applying logarithms to remove // exponentials and make them affine. - rw!("log_eq_log"; "(eq ?a ?b)" => "(eq (log ?a) (log ?b))" + rw!("log_eq_log"; "(eq ?a ?b)" => "(eq (log ?a) (log ?b))" if is_gt_zero("?a") if is_gt_zero("?b")), /* Less than or equal rules. */ rw!("le_sub_iff_add_le"; "(le ?a (sub ?b ?c))" => "(le (add ?a ?c) ?b)"), - + rw!("le_sub_iff_add_le-rev"; "(le (add ?a ?c) ?b)" => "(le ?a (sub ?b ?c))"), rw!("sub_le_iff_le_add"; "(le (sub ?a ?c) ?b)" => "(le ?a (add ?b ?c))"), rw!("sub_le_iff_le_add-rev"; "(le ?a (add ?b ?c))" => "(le (sub ?a ?c) ?b)"), - rw!("div_le_iff"; "(le (div ?a ?c) ?b)" => "(le ?a (mul ?b ?c))" + rw!("div_le_iff"; "(le (div ?a ?c) ?b)" => "(le ?a (mul ?b ?c))" if is_gt_zero("?c")), - - rw!("div_le_iff-rev"; "(le ?a (mul ?b ?c))" => "(le (div ?a ?c) ?b)" + + rw!("div_le_iff-rev"; "(le ?a (mul ?b ?c))" => "(le (div ?a ?c) ?b)" if is_gt_zero("?c")), - rw!("log_le_log"; "(le (log ?a) (log ?b))" => "(le ?a ?b)" - if is_gt_zero("?a") if is_gt_zero("?b")), + rw!("log_le_log"; "(le (log ?a) (log ?b))" => "(le ?a ?b)" + if is_gt_zero("?a") if is_gt_zero("?b")), rw!("log_le_log-rev"; "(le ?a ?b)" => "(le (log ?a) (log ?b))" if is_gt_zero("?a") if is_gt_zero("?b")), - + rw!("pow_two_le_pow_two"; "(le (pow ?a 2) (pow ?b 2))" => "(le ?a ?b)" if is_ge_zero("?a") if is_ge_zero("?b")), - rw!("pow_two_le_pow_two-rev"; "(le ?a ?b)" => "(le (pow ?a 2) (pow ?b 2))" + rw!("pow_two_le_pow_two-rev"; "(le ?a ?b)" => "(le (pow ?a 2) (pow ?b 2))" if is_ge_zero("?a") if is_ge_zero("?b")), @@ -64,10 +64,10 @@ pub fn rules() -> Vec> { vec![ rw!("add_zero"; "(add ?a 0)" => "?a"), - rw!("add_comm"; "(add ?a ?b)" => "(add ?b ?a)"), + rw!("add_comm"; "(add ?a ?b)" => "(add ?b ?a)"), rw!("add_assoc"; "(add (add ?a ?b) ?c)" => "(add ?a (add ?b ?c))"), - + rw!("sub_self"; "(sub ?a ?a)" => "0"), rw!("one_mul"; "(mul 1 ?a)" => "?a"), @@ -95,7 +95,7 @@ pub fn rules() -> Vec> { vec![ rw!("mul_add"; "(mul ?a (add ?b ?c))" => "(add (mul ?a ?b) (mul ?a ?c))"), rw!("mul_add-rev"; "(add (mul ?a ?b) (mul ?a ?c))" => "(mul ?a (add ?b ?c))"), - + rw!("mul_sub"; "(mul ?a (sub ?b ?c))" => "(sub (mul ?a ?b) (mul ?a ?c))"), rw!("mul_sub-rev"; "(sub (mul ?a ?b) (mul ?a ?c))" => "(mul ?a (sub ?b ?c))"), @@ -112,19 +112,19 @@ pub fn rules() -> Vec> { vec![ rw!("div_mul"; "(mul (div ?a ?b) ?c)" => "(div ?a (div ?b ?c))" if is_not_zero("?b") if is_not_zero("?c")), - + rw!("div_mul-rev"; "(div ?a (div ?b ?c))" => "(mul (div ?a ?b) ?c)" if is_not_zero("?b") if is_not_zero("?c")), rw!("mul_div_mul_left"; "(div (mul ?c ?a) (mul ?c ?b))" => "(div ?a ?b)" if is_not_zero("?c")), - rw!("div_div"; "(div (div ?a ?b) ?c)" => "(div ?a (mul ?b ?c))" + rw!("div_div"; "(div (div ?a ?b) ?c)" => "(div ?a (mul ?b ?c))" if is_not_zero("?b") if is_not_zero("?c")), - rw!("div_div-rev"; "(div ?a (mul ?b ?c))" => "(div (div ?a ?b) ?c)" + rw!("div_div-rev"; "(div ?a (mul ?b ?c))" => "(div (div ?a ?b) ?c)" if is_not_zero("?b") if is_not_zero("?c")), - + rw!("div_self"; "(div ?a ?a)" => "1" if is_not_zero("?a")), rw!("inv_eq_one_div"; "(inv ?a)" => "(div 1 ?a)" if is_not_zero("?a")), @@ -150,7 +150,7 @@ pub fn rules() -> Vec> { vec![ rw!("mul_pow-rev"; "(mul (pow ?a ?n) (pow ?b ?n))" => "(pow (mul ?a ?b) ?n)" if is_ge_zero("?a") if is_ge_zero("?b")), - rw!("pow_mul"; "(pow ?a (mul ?n ?m))" => "(pow (pow ?a ?n) ?m)" + rw!("pow_mul"; "(pow ?a (mul ?n ?m))" => "(pow (pow ?a ?n) ?m)" if is_ge_zero("?a")), rw!("pow_mul-rev"; "(pow (pow ?a ?n) ?m)" => "(pow ?a (mul ?n ?m))" @@ -158,11 +158,11 @@ pub fn rules() -> Vec> { vec![ rw!("div_pow"; "(pow (div ?a ?b) ?n)" => "(div (pow ?a ?n) (pow ?b ?n))" if is_ge_zero("?a") if is_gt_zero("?b")), - + rw!("div_pow-rev"; "(div (pow ?a ?n) (pow ?b ?n))" => "(pow (div ?a ?b) ?n)" if is_ge_zero("?a") if is_gt_zero("?b")), - rw!("pow_sub"; "(pow ?a (sub ?b ?c))" => "(div (pow ?a ?b) (pow ?a ?c))" + rw!("pow_sub"; "(pow ?a (sub ?b ?c))" => "(div (pow ?a ?b) (pow ?a ?c))" if is_gt_zero("?a")), rw!("pow_sub-rev"; "(div (pow ?a ?b) (pow ?a ?c))" => "(pow ?a (sub ?b ?c))" @@ -173,8 +173,8 @@ pub fn rules() -> Vec> { vec![ rw!("div_pow_eq_mul_pow_neg-rev"; "(mul ?a (pow ?b (neg ?c)))" => "(div ?a (pow ?b ?c))" if is_gt_zero("?b")), - - rw!("one_div_eq_pow_neg_one"; "(div 1 ?a)" => "(pow ?a (neg 1))" + + rw!("one_div_eq_pow_neg_one"; "(div 1 ?a)" => "(pow ?a (neg 1))" if is_gt_zero("?a")), rw!("sqrt_eq_rpow"; "(sqrt ?a)" => "(pow ?a 0.5)"), @@ -188,7 +188,7 @@ pub fn rules() -> Vec> { vec![ rw!("pow_half_two"; "(pow (pow ?a 0.5) 2)" => "?a" if is_ge_zero("?a")), rw!("pow_half_two-rev"; "?a" => "(pow (pow ?a 0.5) 2)" if is_ge_zero("?a")), - + rw!("binomial_two"; "(pow (add ?a ?b) 2)" => "(add (pow ?a 2) (add (mul 2 (mul ?a ?b)) (pow ?b 2)))"), rw!("inv_eq_pow_neg_one"; "(inv ?a)" => "(pow ?a (neg 1))" if is_not_zero("?a")), @@ -203,7 +203,7 @@ pub fn rules() -> Vec> { vec![ rw!("exp_sub"; "(exp (sub ?a ?b))" => "(div (exp ?a) (exp ?b))"), rw!("exp_sub-rev"; "(div (exp ?a) (exp ?b))" => "(exp (sub ?a ?b))"), - + rw!("exp_mul"; "(exp (mul ?a ?b))" => "(pow (exp ?a) ?b)"), rw!("exp_mul-rev"; "(pow (exp ?a) ?b)" => "(exp (mul ?a ?b))"), @@ -212,33 +212,33 @@ pub fn rules() -> Vec> { vec![ rw!("exp_neg_eq_one_div-rev"; "(div 1 (exp ?a))" => "(exp (neg ?a))"), - rw!("log_mul"; "(log (mul ?a ?b))" => "(add (log ?a) (log ?b))" + rw!("log_mul"; "(log (mul ?a ?b))" => "(add (log ?a) (log ?b))" if is_gt_zero("?a") if is_gt_zero("?b")), - - rw!("log_mul-rev"; "(add (log ?a) (log ?b))" => "(log (mul ?a ?b))" + + rw!("log_mul-rev"; "(add (log ?a) (log ?b))" => "(log (mul ?a ?b))" if is_gt_zero("?a") if is_gt_zero("?b")), - rw!("log_div"; "(log (div ?a ?b))" => "(sub (log ?a) (log ?b))" + rw!("log_div"; "(log (div ?a ?b))" => "(sub (log ?a) (log ?b))" if is_gt_zero("?a") if is_gt_zero("?b")), - + rw!("log_div-rev"; "(sub (log ?a) (log ?b))" => "(log (div ?a ?b))" if is_gt_zero("?a") if is_gt_zero("?b")), - - rw!("log_pow"; "(log (pow ?a ?b))" => "(mul ?b (log ?a))" + + rw!("log_pow"; "(log (pow ?a ?b))" => "(mul ?b (log ?a))" if is_gt_zero("?a")), rw!("log_pow-rev"; "(mul ?b (log ?a))" => "(log (pow ?a ?b))" if is_gt_zero("?a")), - - rw!("log_pow_nat"; "(log (pow ?a ?b))" => "(mul ?b (log ?a))" + + rw!("log_pow_nat"; "(log (pow ?a ?b))" => "(mul ?b (log ?a))" if is_nat("?b")), rw!("log_pow_nat-rev"; "(mul ?b (log ?a))" => "(log (pow ?a ?b))" if is_nat("?b")), rw!("exp_log"; "(exp (log ?a))" => "?a" if is_gt_zero("?a")), - - rw!("log_exp"; "(log (exp ?a))" => "?a"), + + rw!("log_exp"; "(log (exp ?a))" => "?a"), /* Abs rules. */ @@ -250,27 +250,27 @@ pub fn rules() -> Vec> { vec![ /* Atom folding and unfolding rules. */ - rw!("xexp_fold"; "(mul ?a (exp ?a))" => "(xexp ?a)" + rw!("xexp_fold"; "(mul ?a (exp ?a))" => "(xexp ?a)" if is_ge_zero("?a")), - - rw!("xexp_unfold"; "(xexp ?a)" => "(mul ?a (exp ?a))" + + rw!("xexp_unfold"; "(xexp ?a)" => "(mul ?a (exp ?a))" if is_ge_zero("?a")), rw!("entr_fold"; "(neg (mul ?a (log ?a)))" => "(entr ?a)" if is_gt_zero("?a")), - + rw!("entr_unfold"; "(entr ?a)" => "(neg (mul ?a (log ?a)))" if is_gt_zero("?a")), rw!("qol_fold"; "(div (pow ?a 2) ?b)" => "(qol ?a ?b)" if is_gt_zero("?b")), - + rw!("qol_unfold"; "(qol ?a ?b)" => "(div (pow ?a 2) ?b)" if is_gt_zero("?b")), rw!("geo_fold"; "(sqrt (mul ?a ?b))" => "(geo ?a ?b)" if is_gt_zero("?a") if is_gt_zero("?b")), - + rw!("geo_unfold"; "(geo ?a ?b)" => "(sqrt (mul ?a ?b))" if is_gt_zero("?a") if is_gt_zero("?b")), @@ -287,11 +287,11 @@ pub fn rules() -> Vec> { vec![ pub fn simple_example_rules() -> Vec> { vec![ rw!("mul-comm"; "(mul ?a ?b)" => "(mul ?b ?a)"), - rw!("le-mul"; "(le ?a (mul ?b ?c))" => "(le (div ?a ?c) ?b)" + rw!("le-mul"; "(le ?a (mul ?b ?c))" => "(le (div ?a ?c) ?b)" if is_gt_zero("?c")), - - rw!("le-mul-rev"; "(le (div ?a ?c) ?b)" => "(le ?a (mul ?b ?c))" + + rw!("le-mul-rev"; "(le (div ?a ?c) ?b)" => "(le ?a (mul ?b ?c))" if is_gt_zero("?c")), - + rw!("exp_neg_eq_one_div-rev"; "(div 1 (exp ?a))" => "(exp (neg ?a))"), ] } From bab87e63c61c762a8a96e524ddb4899a1af744fd Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 11:31:35 -0500 Subject: [PATCH 148/189] feat: extend positivity with `(a * x) ^ 2 - x ^ 2` --- CvxLean/Tactic/Arith/PositivityExt.lean | 42 +++++++++++++++++++++++-- 1 file changed, 40 insertions(+), 2 deletions(-) diff --git a/CvxLean/Tactic/Arith/PositivityExt.lean b/CvxLean/Tactic/Arith/PositivityExt.lean index 786443ca..d63b3c8a 100644 --- a/CvxLean/Tactic/Arith/PositivityExt.lean +++ b/CvxLean/Tactic/Arith/PositivityExt.lean @@ -19,7 +19,8 @@ lemma Real.one_sub_one_div_sq_nonneg_of_le_one {x : ℝ} : @[positivity ((1 / (_ : ℝ) ^ (2 : ℝ)) - 1)] def evalOneDivSqSubOne : PositivityExt where eval {_ _α} zα pα e := do - let (.app (.app _sub (.app (.app _div _one) (.app (.app _pow (x : Q(ℝ))) _two))) _one') ← + let (.app (.app _sub + (.app (.app _div _one) (.app (.app _pow (x : Q(ℝ))) _two))) _one') ← withReducible (whnf e) | throwError "not ((1 / x ^ 2) - 1)" -- 0 < x ? let h1 := @@ -91,7 +92,8 @@ def evalExpSubOne : PositivityExt where eval {_ _α} zα pα e := do | _ => pure .none -lemma Real.one_sub_div_exp_pos_of_pos {x : ℝ} : 0 < x → 0 < 1 - 1 / Real.exp x := +lemma Real.one_sub_div_exp_pos_of_pos {x : ℝ} : + 0 < x → 0 < 1 - 1 / Real.exp x := fun h => by field_simp; positivity @[positivity (1 - (1 / (Real.exp (_ : ℝ))))] @@ -105,4 +107,40 @@ def evalOneSubDivExp : PositivityExt where eval {_ _α} zα pα e := do | _ => pure .none +lemma Real.scaled_sq_diff_pos_of_pos {a x : ℝ} : + 0 < a - 1 → 0 < x → 0 < (a * x) ^ 2 - x ^ 2 := + fun h1 h2 => by + have ha : 0 ≤ a := by linarith + have ha1 : 1 < a := by linarith + have hx : 0 ≤ x := by linarith + have hx2 : 0 < x ^ 2 := by positivity + rw [sub_pos, Real.mul_rpow ha hx, ←div_lt_iff hx2, div_self (ne_of_gt hx2)] + simp [ha1, abs_eq_self.mpr ha] + +@[positivity (((_ : ℝ) * (_ : ℝ)) ^ (2 : ℝ)) - ((_ : ℝ) ^ (2 : ℝ))] +def evalSubMulSqSq : PositivityExt where eval {_ _α} zα pα e := do + let (.app (.app _sub + (.app (.app _pow (.app (.app _mul (a : Q(ℝ))) (x : Q(ℝ)))) _two)) + (.app (.app _pow' (x' : Q(ℝ))) _two')) ← + withReducible (whnf e) | throwError "not ((a * x) ^ 2 - x ^ 2)" + if !(← isDefEq x x') then + return .none + -- 0 < a - 1 ? + let h1 := ← do + match ← core zα pα (q($a - 1) : Q(ℝ)) with + | .positive pa => return some pa + | _ => return none + -- 0 < x ? + let h2 := ← do + match ← core zα pα x with + | .positive pa => return some pa + | _ => return none + -- If 0 < a - 1 and 0 < x, then 0 < (a * x) ^ 2 - x ^ 2 + match h1, h2 with + | some h1', some h2' => + let pa' ← mkAppM ``Real.scaled_sq_diff_pos_of_pos #[h1', h2'] + pure (.positive pa') + | _, _ => + pure .none + end Mathlib.Meta.Positivity From a77fda2e255d4a11b0f536080a4e8106774bfc53 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 12:28:05 -0500 Subject: [PATCH 149/189] feat: `positivity_maybe_norm_num` --- CvxLean/Tactic/Arith/Arith.lean | 58 ++++++++++++++++++++++++++++++++- 1 file changed, 57 insertions(+), 1 deletion(-) diff --git a/CvxLean/Tactic/Arith/Arith.lean b/CvxLean/Tactic/Arith/Arith.lean index b9138c38..dd06b94b 100644 --- a/CvxLean/Tactic/Arith/Arith.lean +++ b/CvxLean/Tactic/Arith/Arith.lean @@ -70,16 +70,72 @@ elab (name := prepare_positivity) "prepare_positivity" : tactic => do let mvarId' ← preparePositivity mvarId replaceMainGoal [mvarId'] +open Mathlib.Meta.Positivity + +def positivityNoFailure (goal : MVarId) : MetaM (List MVarId) := do + let t : Q(Prop) ← withReducible goal.getType' + try + let p ← solve t + goal.assign p + pure [] + catch _ => + pure [goal] + +elab (name := positivity) "positivity_no_failure" : tactic => do + liftMetaTactic fun g => positivityNoFailure g + end Tactic syntax "positivity!" : tactic macro_rules | `(tactic| positivity!) => - `(tactic| cases_and; prepare_positivity; positivity) + `(tactic| cases_and; prepare_positivity; norm_cast; positivity) syntax "arith" : tactic macro_rules | `(tactic| arith) => `(tactic| (first | linarith | positivity! | norm_num_simp_pow)) + +open Real + +lemma xxx : ∀ (x : ℝ × ℝ × ℝ × ℝ), + log (OfScientific.ofScientific 10 true 0) ≤ + x.2.2.1 - x.1 - + OfScientific.ofScientific 5 true 1 * + log (rexp (OfScientific.ofScientific 2 true 0 * x.1) + rexp (OfScientific.ofScientific 2 true 0 * x.2.1)) → + x.2.1 ≤ + x.2.2.1 - + (OfScientific.ofScientific 5 true 1 * + log + (rexp (OfScientific.ofScientific 2 true 0 * x.1) + + rexp (OfScientific.ofScientific 2 true 0 * x.2.1)) + + log (OfScientific.ofScientific 20 true 0)) → + OfScientific.ofScientific 0 true 0 ≤ x.1 → + x.1 ≤ log (OfScientific.ofScientific 100 true 0) → + OfScientific.ofScientific 0 true 0 ≤ x.2.1 → + x.2.1 ≤ log (OfScientific.ofScientific 100 true 0) → + sqrt (rexp x.2.2.1 / (314159 / 50000) + rexp x.2.2.2 ^ 2) ≤ 10 → + (log + (rexp x.2.2.2 ^ OfScientific.ofScientific 2 true 0 * + ((OfScientific.ofScientific 11 true 0 / OfScientific.ofScientific 10 true 0) ^ + OfScientific.ofScientific 2 true 0 - + OfScientific.ofScientific 1 true 0)) ≤ + log + (rexp x.2.2.1 / + (OfScientific.ofScientific 314159 true 0 / OfScientific.ofScientific 50000 true 0)) ↔ + log (rexp x.2.2.2 ^ OfScientific.ofScientific 2 true 0) + + log + ((OfScientific.ofScientific 11 true 0 / OfScientific.ofScientific 10 true 0) ^ + OfScientific.ofScientific 2 true 0 - + OfScientific.ofScientific 1 true 0) ≤ + log + (rexp x.2.2.1 / + (OfScientific.ofScientific 314159 true 0 / OfScientific.ofScientific 50000 true 0))) := by { + norm_num_clean_up + intros + rw [log_mul] + { positivity! } + { norm_cast; positivity! } + } From d808030b7af2e18a972a995b88188108862a6af5 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 12:28:27 -0500 Subject: [PATCH 150/189] feat: `positivity_maybe_norm_num` in `positivity!` --- CvxLean/Tactic/Arith/Arith.lean | 29 +++++++++++++++++------------ 1 file changed, 17 insertions(+), 12 deletions(-) diff --git a/CvxLean/Tactic/Arith/Arith.lean b/CvxLean/Tactic/Arith/Arith.lean index dd06b94b..39e5ab0d 100644 --- a/CvxLean/Tactic/Arith/Arith.lean +++ b/CvxLean/Tactic/Arith/Arith.lean @@ -72,17 +72,22 @@ elab (name := prepare_positivity) "prepare_positivity" : tactic => do open Mathlib.Meta.Positivity -def positivityNoFailure (goal : MVarId) : MetaM (List MVarId) := do - let t : Q(Prop) ← withReducible goal.getType' - try - let p ← solve t - goal.assign p - pure [] - catch _ => - pure [goal] - -elab (name := positivity) "positivity_no_failure" : tactic => do - liftMetaTactic fun g => positivityNoFailure g +/-- Call `positivity` but if the expression has no free variables, we try to +apply `norm_num` first. -/ +def positivityMaybeNormNum : TacticM Unit := + withMainContext do + let g ← getMainGoal + let t : Q(Prop) ← withReducible g.getType' + let fvs := (collectFVars {} t).fvarSet + let tac ← + if fvs.isEmpty then + `(tactic| (norm_num <;> positivity)) + else + `(tactic| positivity) + let [] ← evalTacticAt tac g | throwError "positivity_maybe_norm_num failed" + +elab (name := positivity) "positivity_maybe_norm_num" : tactic => + positivityMaybeNormNum end Tactic @@ -90,7 +95,7 @@ syntax "positivity!" : tactic macro_rules | `(tactic| positivity!) => - `(tactic| cases_and; prepare_positivity; norm_cast; positivity) + `(tactic| cases_and; prepare_positivity; positivity_maybe_norm_num) syntax "arith" : tactic From 372540633d1c570c059880a9df719d6716294343 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 12:29:09 -0500 Subject: [PATCH 151/189] chore: clean up `convexify` unit tests --- CvxLean/Test/Convexify/Unit.lean | 50 ++++++++------------------------ 1 file changed, 12 insertions(+), 38 deletions(-) diff --git a/CvxLean/Test/Convexify/Unit.lean b/CvxLean/Test/Convexify/Unit.lean index 1219e1a8..043d152e 100644 --- a/CvxLean/Test/Convexify/Unit.lean +++ b/CvxLean/Test/Convexify/Unit.lean @@ -13,7 +13,7 @@ open CvxLean Minimization Real /- Objective function rules. -/ -- map_objFun_log (obj) --- NOTE(RFM): This works because an affine objective is prefered. +-- NOTE: This works because an affine objective is prefered. def mapObjFunLogObj := optimization (x : ℝ) minimize (exp x) @@ -56,7 +56,7 @@ time_cmd equivalence logEqLogConstrRed/logEqLogConstrAuto : logEqLogConstr := by /- Less than or equal rules. -/ -- le_sub_iff_add_le (constr) --- NOTE(RFM): This uses le_sub_iff_add_le because 2 * x is preferred over x + x. +-- NOTE: This uses le_sub_iff_add_le because 2 * x is preferred over x + x. def leSubIffAddLeConstr := optimization (x y : ℝ) minimize (0 : ℝ) @@ -81,7 +81,7 @@ time_cmd equivalence leSubIffAddLeConstrRevRed/leSubIffAddLeConstrRevAuto : leSu #print leSubIffAddLeConstrRevAuto -- sub_le_iff_le_add (constr) --- NOTE(RFM): same reasoning as le_sub_iff_add_le. +-- NOTE: same reasoning as le_sub_iff_add_le. def subLeIffLeAddConstr := optimization (x y : ℝ) minimize (0 : ℝ) @@ -237,7 +237,7 @@ time_cmd equivalence addZeroConstrRed/addZeroConstrAuto : addZeroConstr := by #print addZeroConstrAuto -- add_comm (obj) --- NOTE(RFM): This uses one_mul-rev because 2 * x is preferred over x + x. +-- NOTE: This uses one_mul-rev because 2 * x is preferred over x + x. def addCommObj := optimization (x : ℝ) minimize (x + (1 + x) : ℝ) @@ -250,7 +250,7 @@ time_cmd equivalence addCommObjRed/addCommObjAuto : addCommObj := by #print addCommObjAuto -- add_comm (constr) --- NOTE(RFM): This uses one_mul-rev because 2 * x is preferred over x + x. +-- NOTE: This uses one_mul-rev because 2 * x is preferred over x + x. def addCommConstr := optimization (x : ℝ) minimize (0 : ℝ) @@ -264,7 +264,7 @@ time_cmd equivalence addCommConstrRed/addCommConstrAuto : addCommConstr := by #print addCommConstrAuto -- add_assoc (obj) --- NOTE(RFM): This uses one_mul-rev because 2 * x is preferred over x + x. +-- NOTE: This uses one_mul-rev because 2 * x is preferred over x + x. def addAssocObj := optimization (x : ℝ) minimize (x + (x + 1) : ℝ) @@ -277,7 +277,7 @@ time_cmd equivalence addAssocObjRed/addAssocObjAuto : addAssocObj := by #print addAssocObjAuto -- add_assoc (constr) --- NOTE(RFM): This uses one_mul-rev because 2 * x is preferred over x + x. +-- NOTE: This uses one_mul-rev because 2 * x is preferred over x + x. def addAssocConstr := optimization (x : ℝ) minimize (0 : ℝ) @@ -338,7 +338,7 @@ time_cmd equivalence oneMulConstrRed/oneMulConstrAuto : oneMulConstr := by #print oneMulConstrAuto -- one_mul-rev (obj) --- NOTE(RFM): This uses one_mul-rev because 2 * x is preferred over x + x. +-- NOTE: This uses one_mul-rev because 2 * x is preferred over x + x. def oneMulRevObj := optimization (x : ℝ) minimize (x + x : ℝ) @@ -351,7 +351,7 @@ time_cmd equivalence oneMulRevObjRed/oneMulRevObjAuto : oneMulRevObj := by #print oneMulRevObjAuto -- one_mul-rev (constr) --- NOTE(RFM): This uses one_mul-rev because 2 * x is preferred over x + x. +-- NOTE: This uses one_mul-rev because 2 * x is preferred over x + x. def oneMulRevConstr := optimization (x : ℝ) minimize (0 : ℝ) @@ -388,7 +388,7 @@ time_cmd equivalence mulZeroConstrRed/mulZeroConstrAuto : mulZeroConstr := by #print mulZeroConstrAuto -- mul_comm (obj) --- NOTE(RFM): Empty domain here. Why? +-- NOTE: Empty domain here. Why? def mulCommObj := optimization (x y : ℝ) minimize (x * (y * (1 / x)) : ℝ) @@ -707,7 +707,6 @@ def addDivConstr := hy : 0 < y h : (exp x + 1) / (exp y) ≤ 1 --- NOTE(RFM): This relies on the extension of positivity above. time_cmd equivalence addDivConstrRed/addDivConstrAuto : addDivConstr := by convexify @@ -815,12 +814,6 @@ time_cmd equivalence divSelfConstrRed/divSelfConstrAuto : divSelfConstr := by /- Power and square root rules. -/ --- pow_add (obj) --- TODO(RFM): Implement power atom. - --- pow_add (constr) --- TODO(RFM): Implement power atom. - -- pow_add-rev (obj) def powAddRevObj := optimization (x : ℝ) @@ -896,12 +889,6 @@ time_cmd equivalence mulPowRevConstrRed/mulPowRevConstrAuto : mulPowRevConstr := #print mulPowRevConstrAuto --- pow_mul (obj) --- TODO(RFM): Implement power atom. - --- pow_mul (constr) --- TODO(RFM): Implement power atom. - -- pow_mul-rev (obj) def powMulRevObj := optimization (x : ℝ) @@ -961,9 +948,8 @@ def divPowRevObj := time_cmd equivalence divPowRevObjRed/divPowRevObjAuto : divPowRevObj := by convexify --- TODO(RFM): does not rewrite the correct term. --- #print divPowRevObjAuto +#print divPowRevObjAuto -- div_pow-rev (constr) def divPowRevConstr := @@ -976,13 +962,7 @@ def divPowRevConstr := time_cmd equivalence divPowRevConstrRed/divPowRevConstrAuto : divPowRevConstr := by convexify --- #print divPowRevConstrAuto - --- pow_sub (obj) --- TODO(RFM): Implement power atom. - --- pow_sub (constr) --- TODO(RFM): Implement power atom. +#print divPowRevConstrAuto -- pow_sub-rev (obj) def powSubRevObj := @@ -1034,12 +1014,6 @@ time_cmd equivalence divPowEqMulPowNegConstrRed/divPowEqMulPowNegConstrAuto : di #print divPowEqMulPowNegConstrAuto --- div_pow_eq_mul_pow_neg-rev (obj) --- TODO(RFM): Implement power atom. - --- div_pow_eq_mul_pow_neg-rev (constr) --- TODO(RFM): Implement power atom. - -- one_div_eq_pow_neg_one (obj) def oneDivEqPowNegOneObj := optimization (x : ℝ) From b4f03938c54d253934c69235176f4ec71bc4faaf Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 12:48:48 -0500 Subject: [PATCH 152/189] fix: type argument order in rewrite wrappers --- CvxLean/Lib/Minimization.lean | 52 +++++++++++++++++------------------ 1 file changed, 26 insertions(+), 26 deletions(-) diff --git a/CvxLean/Lib/Minimization.lean b/CvxLean/Lib/Minimization.lean index 221bd8e1..d22c5960 100644 --- a/CvxLean/Lib/Minimization.lean +++ b/CvxLean/Lib/Minimization.lean @@ -2,29 +2,29 @@ import Mathlib.Order.Bounds.Basic import Mathlib.Data.Set.Function /-- Type of an optimization problem. -/ -structure Minimization (D : Type) (R : Type) where +structure Minimization (D : Type) (R : Type) where objFun : D → R constraints : D → Prop namespace Minimization -variable {D E R S : Type} [Preorder R] [Preorder S] +variable {D E R S : Type} [Preorder R] [Preorder S] variable (p : Minimization D R) (q : Minimization E R) /-- A feasible point is a point in the domain that satisfies the constraints. -/ -structure FeasPoint where - point : D +structure FeasPoint where + point : D feasibility : p.constraints point /-- A solution is a feasible point that is also optimal. -/ -structure Solution where - point : D - feasibility : p.constraints point +structure Solution where + point : D + feasibility : p.constraints point optimality : ∀ y : p.FeasPoint, p.objFun point ≤ p.objFun y.point section Reductions -/-- Reduction from solution to solution maintaining equivalence through +/-- Reduction from solution to solution maintaining equivalence through functions `f` and `g`. For most purposes, this might be enough. -/ def simple_reduction (sol : q.Solution) @@ -49,11 +49,11 @@ def simple_reduction /-- Decompose constraint by adding another equality constraint. -/ def decompose_constraint (g : D → E) (c : D → E → Prop) - (hc : ∀ x, p.constraints x = c x (g x)) + (hc : ∀ x, p.constraints x = c x (g x)) (sol : Solution - { objFun := fun (y : E × D) => objFun p y.snd, + { objFun := fun (y : E × D) => objFun p y.snd, constraints := fun (y : E × D) => y.fst = g y.snd ∧ c y.snd y.fst }) : - p.Solution := + p.Solution := simple_reduction p _ sol (fun x => (g x, x)) (fun x => x.2) (fun {x} hx => le_refl _) @@ -65,8 +65,8 @@ simple_reduction p _ sol def eq_to_le_left (e: Equiv D (S × E)) (f : E → S) (c : D → Prop) (hc : ∀ {x}, p.constraints x ↔ ((e.toFun x).1 = f (e.toFun x).2 ∧ c x)) - (h_objFun : ∀ x r s, p.objFun (e.symm.toFun (r,x)) = p.objFun (e.symm.toFun (s,x))) - (h_mono: ∀ x r s, r ≤ s → c (e.symm.toFun (r, x)) → c (e.symm.toFun (s, x))) + (h_objFun : ∀ x r s, p.objFun (e.symm.toFun (r,x)) = p.objFun (e.symm.toFun (s,x))) + (h_mono: ∀ x r s, r ≤ s → c (e.symm.toFun (r, x)) → c (e.symm.toFun (s, x))) (sol : Solution { objFun := p.objFun, constraints := fun x => (e.toFun x).1 ≤ f (e.toFun x).2 ∧ c x } ) : @@ -86,11 +86,11 @@ simple_reduction p _ sol simp_all ) /-- -/ -def eq_to_le_right +def eq_to_le_right (e: Equiv D (S × E)) (f : E → S) (c : D → Prop) (hc : ∀ {x}, p.constraints x ↔ (f (e.toFun x).2 = (e.toFun x).1 ∧ c x)) - (h_objFun : ∀ x r s, p.objFun (e.symm.toFun ⟨r, x⟩) = p.objFun (e.symm.toFun ⟨s, x⟩)) - (h_mono: ∀ x r s, r ≤ s → c (e.symm.toFun (s, x)) → c (e.symm.toFun ⟨r, x⟩)) + (h_objFun : ∀ x r s, p.objFun (e.symm.toFun ⟨r, x⟩) = p.objFun (e.symm.toFun ⟨s, x⟩)) + (h_mono: ∀ x r s, r ≤ s → c (e.symm.toFun (s, x)) → c (e.symm.toFun ⟨r, x⟩)) (sol : Solution { objFun := p.objFun, constraints := fun x => f (e.toFun x).2 ≤ (e.toFun x).1 ∧ c x }) : @@ -121,7 +121,7 @@ def linearization_mono {of : D → R} {cs : D → Prop} (sol : Solution { objFun := fun (y : S × D) => f y.1 y.2, constraints := fun y => y.1 ≤ g y.2 ∧ c y.1 y.2 }) : - Solution {objFun := of, constraints := cs} := + Solution {objFun := of, constraints := cs} := simple_reduction _ _ sol (fun x => (g x, x)) (fun x => x.2) (fun {x} hx => le_of_eq (hof _).symm) @@ -141,7 +141,7 @@ def linearization_antimono {of : D → R} {cs : D → Prop} (sol : Solution { objFun := fun (y : S × D) => f y.1 y.2, constraints := fun y => g y.2 ≤ y.1 ∧ c y.1 y.2 }) : - Solution {objFun := of, constraints := cs} := + Solution {objFun := of, constraints := cs} := simple_reduction _ _ sol (fun x => (g x, x)) (fun x => x.2) (fun {x} hx => le_of_eq (hof _).symm) @@ -203,14 +203,14 @@ def graph_expansion_least_forall {of : D → R} {cs : D → Prop} { objFun := fun (y : (I → S) × D) => of y.2, constraints := fun (y : (I → S) × D) => (∀ i, d (y.1 i) y.2) ∧ (∀ i, c (y.1 i) y.2) }) : Solution {objFun := of, constraints := cs} := - @graph_expansion_least D R _ of cs (I → S) g + @graph_expansion_least D R _ of cs (I → S) g (fun y x => ∀ i, c (y i) x) (fun y x => ∀ i, d (y i) x) (fun y x => of x) - ⟨fun a i => le_refl (a i), + ⟨fun a i => le_refl (a i), fun a b c hab hbc i => le_trans (hab i) (hbc i), fun a b => Iff.refl _⟩ - (fun x v hc => ⟨fun i => (hg x (v i) i (hc i)).1, + (fun x v hc => ⟨fun i => (hg x (v i) i (hc i)).1, fun v' c i => (hg x (v i) i (hc i)).2 (c i)⟩) (fun x => rfl) hcs @@ -270,7 +270,7 @@ simple_reduction _ _ sol f g /- Rewrites used in `convexify` under the `reduction` command. -/ section Rewrites -def rewrite_objective {D R} [Preorder R] {f g : D → R} {cs : D → Prop} +def rewrite_objective {D R} [Preorder R] {f g : D → R} {cs : D → Prop} (hrw : ∀ x, cs x → f x = g x) (sol : Solution { objFun := g, constraints := cs }) : Solution { objFun := f, constraints := cs } := @@ -427,7 +427,7 @@ def rewrite_constraint_7_last {D R} [Preorder R] {c1 c2 c3 c4 c5 c6 c7 c7' : D (fun {x} hx => by simp only [hrw x hx.1 hx.2.1 hx.2.2.1 hx.2.2.2.1 hx.2.2.2.2.1 hx.2.2.2.2.2.1] at hx; exact hx) (fun {x} hx => by simp only [←hrw x hx.1 hx.2.1 hx.2.2.1 hx.2.2.2.1 hx.2.2.2.2.1 hx.2.2.2.2.2.1] at hx; exact hx) -def rewrite_constraint_8 {R D} [Preorder R] {c1 c2 c3 c4 c5 c6 c7 c8 c8' : D → Prop} {cs : D → Prop} {f : D → R} +def rewrite_constraint_8 {D R} [Preorder R] {c1 c2 c3 c4 c5 c6 c7 c8 c8' : D → Prop} {cs : D → Prop} {f : D → R} (hrw : ∀ x, c1 x → c2 x → c3 x → c4 x → c5 x → c6 x → c7 x → cs x → (c8 x ↔ c8' x)) (sol : Solution { objFun := f, constraints := fun x => c1 x ∧ c2 x ∧ c3 x ∧ c4 x ∧ c5 x ∧ c6 x ∧ c7 x ∧ c8' x ∧ cs x }) : Solution { objFun := f, constraints := fun x => c1 x ∧ c2 x ∧ c3 x ∧ c4 x ∧ c5 x ∧ c6 x ∧ c7 x ∧ c8 x ∧ cs x } := @@ -437,7 +437,7 @@ def rewrite_constraint_8 {R D} [Preorder R] {c1 c2 c3 c4 c5 c6 c7 c8 c8' : D → (fun {x} hx => by simp only [hrw x hx.1 hx.2.1 hx.2.2.1 hx.2.2.2.1 hx.2.2.2.2.1 hx.2.2.2.2.2.1 hx.2.2.2.2.2.2.1 hx.2.2.2.2.2.2.2.2] at hx; exact hx) (fun {x} hx => by simp only [←hrw x hx.1 hx.2.1 hx.2.2.1 hx.2.2.2.1 hx.2.2.2.2.1 hx.2.2.2.2.2.1 hx.2.2.2.2.2.2.1 hx.2.2.2.2.2.2.2.2] at hx; exact hx) -def rewrite_constraint_8_last {R D} [Preorder R] {c1 c2 c3 c4 c5 c6 c7 c8 c8' : D → Prop} {f : D → R} +def rewrite_constraint_8_last {D R} [Preorder R] {c1 c2 c3 c4 c5 c6 c7 c8 c8' : D → Prop} {f : D → R} (hrw : ∀ x, c1 x → c2 x → c3 x → c4 x → c5 x → c6 x → c7 x → (c8 x ↔ c8' x)) (sol : Solution { objFun := f, constraints := fun x => c1 x ∧ c2 x ∧ c3 x ∧ c4 x ∧ c5 x ∧ c6 x ∧ c7 x ∧ c8' x }) : Solution { objFun := f, constraints := fun x => c1 x ∧ c2 x ∧ c3 x ∧ c4 x ∧ c5 x ∧ c6 x ∧ c7 x ∧ c8 x } := @@ -457,7 +457,7 @@ def rewrite_constraint_9 {D R} [Preorder R] {c1 c2 c3 c4 c5 c6 c7 c8 c9 c9' : D (fun {x} hx => by simp only [hrw x hx.1 hx.2.1 hx.2.2.1 hx.2.2.2.1 hx.2.2.2.2.1 hx.2.2.2.2.2.1 hx.2.2.2.2.2.2.1 hx.2.2.2.2.2.2.2.1 hx.2.2.2.2.2.2.2.2.2] at hx; exact hx) (fun {x} hx => by simp only [←hrw x hx.1 hx.2.1 hx.2.2.1 hx.2.2.2.1 hx.2.2.2.2.1 hx.2.2.2.2.2.1 hx.2.2.2.2.2.2.1 hx.2.2.2.2.2.2.2.1 hx.2.2.2.2.2.2.2.2.2] at hx; exact hx) -def rewrite_constraint_9_last {R D} [Preorder R] {c1 c2 c3 c4 c5 c6 c7 c8 c9 c9' : D → Prop} {f : D → R} +def rewrite_constraint_9_last {D R} [Preorder R] {c1 c2 c3 c4 c5 c6 c7 c8 c9 c9' : D → Prop} {f : D → R} (hrw : ∀ x, c1 x → c2 x → c3 x → c4 x → c5 x → c6 x → c7 x → c8 x → (c9 x ↔ c9' x)) (sol : Solution { objFun := f, constraints := fun x => c1 x ∧ c2 x ∧ c3 x ∧ c4 x ∧ c5 x ∧ c6 x ∧ c7 x ∧ c8 x ∧ c9' x }) : Solution { objFun := f, constraints := fun x => c1 x ∧ c2 x ∧ c3 x ∧ c4 x ∧ c5 x ∧ c6 x ∧ c7 x ∧ c8 x ∧ c9 x } := @@ -477,7 +477,7 @@ def rewrite_constraint_10 {D R} [Preorder R] {c1 c2 c3 c4 c5 c6 c7 c8 c9 c10 c10 (fun {x} hx => by simp only [hrw x hx.1 hx.2.1 hx.2.2.1 hx.2.2.2.1 hx.2.2.2.2.1 hx.2.2.2.2.2.1 hx.2.2.2.2.2.2.1 hx.2.2.2.2.2.2.2.1 hx.2.2.2.2.2.2.2.2.1 hx.2.2.2.2.2.2.2.2.2.2] at hx; exact hx) (fun {x} hx => by simp only [←hrw x hx.1 hx.2.1 hx.2.2.1 hx.2.2.2.1 hx.2.2.2.2.1 hx.2.2.2.2.2.1 hx.2.2.2.2.2.2.1 hx.2.2.2.2.2.2.2.1 hx.2.2.2.2.2.2.2.2.1 hx.2.2.2.2.2.2.2.2.2.2] at hx; exact hx) -def rewrite_constraint_10_last {R D} [Preorder R] {c1 c2 c3 c4 c5 c6 c7 c8 c9 c10 c10' : D → Prop} {f : D → R} +def rewrite_constraint_10_last {D R} [Preorder R] {c1 c2 c3 c4 c5 c6 c7 c8 c9 c10 c10' : D → Prop} {f : D → R} (hrw : ∀ x, c1 x → c2 x → c3 x → c4 x → c5 x → c6 x → c7 x → c8 x → c9 x → (c10 x ↔ c10' x)) (sol : Solution { objFun := f, constraints := fun x => c1 x ∧ c2 x ∧ c3 x ∧ c4 x ∧ c5 x ∧ c6 x ∧ c7 x ∧ c8 x ∧ c9 x ∧ c10' x }) : Solution { objFun := f, constraints := fun x => c1 x ∧ c2 x ∧ c3 x ∧ c4 x ∧ c5 x ∧ c6 x ∧ c7 x ∧ c8 x ∧ c9 x ∧ c10 x } := From 0babbfeeab1eb924fb697edf289a2882bc2def56 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 13:04:58 -0500 Subject: [PATCH 153/189] fix: weaken `egg`'s DCP check --- CvxLean/Test/Convexify/DGP.lean | 31 +++++++++++++++++++++++++++++++ 1 file changed, 31 insertions(+) diff --git a/CvxLean/Test/Convexify/DGP.lean b/CvxLean/Test/Convexify/DGP.lean index 9a005577..9294f454 100644 --- a/CvxLean/Test/Convexify/DGP.lean +++ b/CvxLean/Test/Convexify/DGP.lean @@ -192,6 +192,7 @@ end GP6 section GP7 -- NOTE: We don't have the power atom yet. +-- TODO: add atoms for ^(1/2) and ^(-1/2). -- objFun : (x ^ (-1)) * y ^ (-1 / 2) * z ^ (-1) + 2.3 * x * z + 4 * x * y * z -- h4 : (1 / 3) * x ^ (-2) * y ^ (-2) + (4 / 3) * y ^ (1 / 2) * z ^ (-1) ≤ 1 def gp7 := @@ -223,6 +224,36 @@ solve dcp7 end GP7 +/- In https://web.stanford.edu/~boyd/papers/pdf/gp_tutorial.pdf section 6.1. -/ +section GP8 + +-- hmin = wmin = 1, hmax = wmax = 100, Rmax = 10, σ = 0.5, π ≈ 3.14159. +def gp8 := + optimization (h w A r : ℝ) + minimize 2 * A * sqrt (w ^ 2 + h ^ 2) + subject to + h1 : 0 < h + h2 : 0 < w + h3 : 0 < A + h4 : 0 < r + h5 : 10 * sqrt (w ^ 2 + h ^ 2) / 2 * h ≤ 0.5 * A + h6 : 20 * sqrt (w ^ 2 + h ^ 2) / 2 * w ≤ 0.5 * A + h7 : 1 ≤ h + h8 : h ≤ 100 + h9 : 1 ≤ w + h10 : w ≤ 100 + h11 : 1.1 * r ≤ sqrt (A / (2 * 3.14159) + r ^ 2) + h12 : sqrt (A / (2 * 3.14159) + r ^ 2) ≤ 10 + +set_option maxHeartbeats 1000000 in +time_cmd reduction red8/dcp8 : gp8 := by + map_exp + convexify + +solve dcp8 + +end GP8 + end end GP From e2273fe09ea0f75af04e79aebc58be87875f335f Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 13:05:28 -0500 Subject: [PATCH 154/189] feat: truss design working --- egg-convexify/src/cost.rs | 65 +++++++++++++++++++++------------------ 1 file changed, 35 insertions(+), 30 deletions(-) diff --git a/egg-convexify/src/cost.rs b/egg-convexify/src/cost.rs index d9e64fba..56859db1 100644 --- a/egg-convexify/src/cost.rs +++ b/egg-convexify/src/cost.rs @@ -245,16 +245,18 @@ impl<'a> CostFunction for DCPCost<'a> { term_size = 1 + get_term_size!(a) + get_term_size!(b); } Optimization::QOL([a, b]) => { - let da_nonneg = domain::option_is_nonneg(get_domain(a).as_ref()); - let da_nonpos = domain::option_is_nonpos(get_domain(a).as_ref()); - let curvature_num = - if da_nonneg { - curvature::of_convex_nondecreasing_fn(get_curvature!(a)) - } else if da_nonpos { - curvature::of_convex_nonincreasing_fn(get_curvature!(a)) - } else { - curvature::of_convex_none_fn(get_curvature!(a)) - }; + // let da_nonneg = domain::option_is_nonneg(get_domain(a).as_ref()); + // let da_nonpos = domain::option_is_nonpos(get_domain(a).as_ref()); + // let curvature_num = + // if da_nonneg { + // curvature::of_convex_nondecreasing_fn(get_curvature!(a)) + // } else if da_nonpos { + // curvature::of_convex_nonincreasing_fn(get_curvature!(a)) + // } else { + // curvature::of_convex_none_fn(get_curvature!(a)) + // }; + // NOTE: Aligned with the atom library. + let curvature_num = curvature::of_convex_none_fn(get_curvature!(a)); let db_pos = domain::option_is_pos(get_domain(b).as_ref()); if db_pos { let curvature_den = curvature::of_convex_nonincreasing_fn(get_curvature!(b)); @@ -282,26 +284,29 @@ impl<'a> CostFunction for DCPCost<'a> { term_size = 1 + get_term_size!(a) + get_term_size!(b); } Optimization::Norm2([a, b]) => { - let da_nonneg = domain::option_is_nonneg(get_domain(a).as_ref()); - let da_nonpos = domain::option_is_nonpos(get_domain(a).as_ref()); - let curvature_a = - if da_nonneg { - curvature::of_convex_nondecreasing_fn(get_curvature!(a)) - } else if da_nonpos { - curvature::of_convex_nonincreasing_fn(get_curvature!(a)) - } else { - curvature::of_convex_none_fn(get_curvature!(a)) - }; - let db_nonneg = domain::option_is_nonneg(get_domain(b).as_ref()); - let db_nonpos = domain::option_is_nonpos(get_domain(b).as_ref()); - let curvature_b = - if db_nonneg { - curvature::of_convex_nondecreasing_fn(get_curvature!(b)) - } else if db_nonpos { - curvature::of_convex_nonincreasing_fn(get_curvature!(b)) - } else { - curvature::of_convex_none_fn(get_curvature!(b)) - }; + // let da_nonneg = domain::option_is_nonneg(get_domain(a).as_ref()); + // let da_nonpos = domain::option_is_nonpos(get_domain(a).as_ref()); + // let curvature_a = + // if da_nonneg { + // curvature::of_convex_nondecreasing_fn(get_curvature!(a)) + // } else if da_nonpos { + // curvature::of_convex_nonincreasing_fn(get_curvature!(a)) + // } else { + // curvature::of_convex_none_fn(get_curvature!(a)) + // }; + // let db_nonneg = domain::option_is_nonneg(get_domain(b).as_ref()); + // let db_nonpos = domain::option_is_nonpos(get_domain(b).as_ref()); + // let curvature_b = + // if db_nonneg { + // curvature::of_convex_nondecreasing_fn(get_curvature!(b)) + // } else if db_nonpos { + // curvature::of_convex_nonincreasing_fn(get_curvature!(b)) + // } else { + // curvature::of_convex_none_fn(get_curvature!(b)) + // }; + // NOTE: Aligned with the atom library. + let curvature_a = curvature::of_convex_none_fn(get_curvature!(a)); + let curvature_b = curvature::of_convex_none_fn(get_curvature!(b)); curvature = curvature::join(curvature_a, curvature_b); num_vars = get_num_vars!(a) + get_num_vars!(b); term_size = 1 + get_term_size!(a) + get_term_size!(b); From c779ebfd98a2d253c52bb5f7023de3b0f3ba26ea Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 13:06:11 -0500 Subject: [PATCH 155/189] chore: move change of variables test --- CvxLean/Test/{Convexify => Other}/ChangeOfVariables.lean | 0 1 file changed, 0 insertions(+), 0 deletions(-) rename CvxLean/Test/{Convexify => Other}/ChangeOfVariables.lean (100%) diff --git a/CvxLean/Test/Convexify/ChangeOfVariables.lean b/CvxLean/Test/Other/ChangeOfVariables.lean similarity index 100% rename from CvxLean/Test/Convexify/ChangeOfVariables.lean rename to CvxLean/Test/Other/ChangeOfVariables.lean From 586ff3fb76009f8c53fe0b41c1323eb40947cc65 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 13:13:26 -0500 Subject: [PATCH 156/189] fix: recover HSD, `10e-6` trick --- CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean | 20 +++++++++---------- .../DCP/AtomLibrary/Fns/QuadOverLin.lean | 4 ++-- CvxLean/Test/Convexify/DQCP.lean | 2 +- 3 files changed, 13 insertions(+), 13 deletions(-) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean index 307d6f38..fb84f360 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean @@ -15,12 +15,12 @@ open Real declare_atom klDiv [convex] (x : ℝ)? (y : ℝ)? : x * log (x / y) - x + y := vconditions (hx : 0 ≤ x) - (hy : 10e-6 ≤ y) + (hy : 1 / 100000 ≤ y) implementationVars (t : ℝ) implementationObjective (y - x - t) implementationConstraints (c1 : expCone t x y) - (c2 : 10e-6 ≤ y) + (c2 : 1 / 100000 ≤ y) solution (t := -(x * log (x / y))) solutionEqualsAtom by ring @@ -73,12 +73,12 @@ vconditionElimination declare_atom klDiv2 [convex] (x : ℝ)? (y : ℝ)? : klDiv x y := vconditions (hx : 0 ≤ x) - (hy : 10e-6 ≤ y) + (hy : 1 / 100000 ≤ y) implementationVars (t : ℝ) implementationObjective (y - x - t) implementationConstraints (c1 : expCone t x y) - (c2 : 10e-6 ≤ y) + (c2 : 1 / 100000 ≤ y) solution (t := -(x * log (x / y))) solutionEqualsAtom by unfold klDiv @@ -108,12 +108,12 @@ declare_atom Vec.klDiv [convex] (m : Nat)& (x : Fin m → ℝ)? (y : Fin m → Vec.klDiv x y := vconditions (hx : 0 ≤ x) - (hy : (fun _ => 10e-6) ≤ y) + (hy : (fun _ => 1 / 100000) ≤ y) implementationVars (t : Fin m → ℝ) implementationObjective (y - x - t) implementationConstraints (c1 : Vec.expCone t x y) - (c2 : Vec.const m 10e-6 ≤ y) + (c2 : Vec.const m (1 / 100000) ≤ y) solution (t := fun i => -((x i) * log ((x i) / (y i)))) solutionEqualsAtom by simp [Vec.klDiv, klDiv]; ext i; simp; ring @@ -128,22 +128,22 @@ feasibility intros i have h := klDiv.feasibility1 (x i) (y i) (hx i) (hy i) unfold posOrthCone at h - simpa using h) + simpa [Vec.const] using h) optimality fun i => by unfold Vec.expCone at c1 apply klDiv.optimality (x i) (y i) (t i) (c1 i) unfold posOrthCone expCone at * - simpa using (c2 i) + simpa [Vec.const] using (c2 i) vconditionElimination (hx : fun i => by unfold Vec.expCone at c1 apply klDiv.vcondElim0 (x i) (y i) (t i) (c1 i) unfold posOrthCone expCone at * - simpa using (c2 i)) + simpa [Vec.const] using (c2 i)) (hy : fun i => by unfold Vec.expCone at c1 apply klDiv.vcondElim1 (x i) (y i) (t i) (c1 i) unfold posOrthCone expCone at * - simpa using (c2 i)) + simpa [Vec.const] using (c2 i)) end CvxLean diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean index 58b05b6f..ff3276eb 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean @@ -17,13 +17,13 @@ open Real -- TODO(RFM): how to do strict positivity and make vcondelimination work? declare_atom quadOverLin [convex] (x : ℝ)? (y : ℝ)- : x ^ 2 / y := vconditions - (hy : 10e-6 ≤ y) + (hy : 1 / 100000 ≤ y) implementationVars (t : ℝ) (y' : ℝ) implementationObjective (t) implementationConstraints (c1 : soCone (y + t) ![y - t, 2 * x]) (c2 : 0 ≤ t) - (c3 : 10e-6 ≤ y) + (c3 : 1 / 100000 ≤ y) solution (t := x ^ 2 / y) (y' := log y) solutionEqualsAtom by rfl feasibility diff --git a/CvxLean/Test/Convexify/DQCP.lean b/CvxLean/Test/Convexify/DQCP.lean index 5aaf3628..87a1f71c 100644 --- a/CvxLean/Test/Convexify/DQCP.lean +++ b/CvxLean/Test/Convexify/DQCP.lean @@ -47,7 +47,7 @@ def hypersonicShapeDesign (a b : ℝ) := optimization (Δx : ℝ) minimize sqrt (1 / (Δx ^ 2) - 1) subject to - h1 : 0 < Δx + h1 : 10e-6 ≤ Δx h2 : Δx ≤ 1 h3 : a * (1 / Δx) - (1 - b) * sqrt (1 - Δx ^ 2) ≤ 0 From c5ef5cf3899a0cde4a75acb5d8f626ab89d04904 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 13:23:26 -0500 Subject: [PATCH 157/189] feat: `EggTree` size --- CvxLean/Tactic/Convexify/Convexify.lean | 2 ++ CvxLean/Tactic/Convexify/Egg/FromMinimization.lean | 5 +++++ 2 files changed, 7 insertions(+) diff --git a/CvxLean/Tactic/Convexify/Convexify.lean b/CvxLean/Tactic/Convexify/Convexify.lean index fefefbf1..76ad52ee 100644 --- a/CvxLean/Tactic/Convexify/Convexify.lean +++ b/CvxLean/Tactic/Convexify/Convexify.lean @@ -285,6 +285,8 @@ def evalConvexify : Tactic := fun stx => match stx with let diff := after - before dbg_trace s!"Egg time: {diff} ms." dbg_trace s!"Number of steps: {steps.size}." + let size := (gStr.map fun (_, t) => t.size).fold 0 Nat.add + dbg_trace s!"Term size: {size}." -- Apply steps. let mut g ← getMainGoal diff --git a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean index 5b7de02a..9814685b 100644 --- a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean +++ b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean @@ -15,6 +15,11 @@ partial def _root_.EggTree.toEggString : Tree String String → String "(" ++ n ++ " " ++ (" ".intercalate childrenStr) ++ ")" | Tree.leaf n => n +/-- Size of the AST. -/ +partial def _root_.EggTree.size : EggTree → Nat + | Tree.node _ children => 1 + (children.map EggTree.size).foldl Nat.add 0 + | Tree.leaf _ => 1 + /-- -/ def _root_.EggTree.ofOCTree (ocTree : OC (String × Tree String String)) : Tree String String := From b939fb3f19ed4c834763a92b0278011940b79804 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 13:36:13 -0500 Subject: [PATCH 158/189] fix: DCP tests --- CvxLean/Test/DCP/Dcp.lean | 7 +++---- 1 file changed, 3 insertions(+), 4 deletions(-) diff --git a/CvxLean/Test/DCP/Dcp.lean b/CvxLean/Test/DCP/Dcp.lean index cadf9de0..a283a2e3 100644 --- a/CvxLean/Test/DCP/Dcp.lean +++ b/CvxLean/Test/DCP/Dcp.lean @@ -22,7 +22,7 @@ noncomputable def test004 : Solution $ minimize exp v subject to cv : 0 ≤ v - cw : 0 < w + cw : 1 / 100000 ≤ w c0 : klDiv v w ≤ 1 := by dcp @@ -112,8 +112,7 @@ optimization (x y : Fin n → ℝ) minimize (0 : ℝ) subject to cx : 0 ≤ x - cy : StrongLT 0 y - -- cy : ∀ i, 0 < (y) i + cy : 1 / 100000 ≤ y c0 : Vec.klDiv x y ≤ x := by dcp @@ -198,7 +197,7 @@ noncomputable def testQuadOverLin : Solution $ minimize x ^ 2 / y subject to c0 : 0 ≤ x - c0 : 0.001 ≤ y + c0 : 1 / 100000 ≤ y := by dcp sorry From ed169323230f3d702a5f5717e951f6c0e7106572 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 15:57:58 -0500 Subject: [PATCH 159/189] chore: ignore evaluation --- .gitignore | 1 + 1 file changed, 1 insertion(+) diff --git a/.gitignore b/.gitignore index 08378766..d15ee674 100644 --- a/.gitignore +++ b/.gitignore @@ -1,6 +1,7 @@ .DS_Store /build +/evaluation /solver/*.cbf /solver/*.sol /egg-convexify/*.dot From b0890778adfcf4cd6e7f96a04ee814f9608c4325 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 16:55:46 -0500 Subject: [PATCH 160/189] fix: panic in `vcond` inference --- CvxLean/Tactic/DCP/Dcp.lean | 51 ++++++++++++++++--------------------- 1 file changed, 22 insertions(+), 29 deletions(-) diff --git a/CvxLean/Tactic/DCP/Dcp.lean b/CvxLean/Tactic/DCP/Dcp.lean index e533fcbd..42f33536 100644 --- a/CvxLean/Tactic/DCP/Dcp.lean +++ b/CvxLean/Tactic/DCP/Dcp.lean @@ -255,25 +255,19 @@ partial def findVConditions (originalConstrVars : Array LocalDecl) (constraints match ← constraints.findIdxM? (isDefEq vcond) with | some i => do vcondData := vcondData.push (Sum.inl i) | none => - -- TODO(RFM): Same issue with background conditions. Find a less hacky way? -- Infer vconditions from constraints. - let vcondProofTyBody ← liftM <| constraints.foldrM mkArrow vcond - let vcondProofTy ← mkForallFVars args vcondProofTyBody - - let (e, _) ← Lean.Elab.Term.TermElabM.run <| Lean.Elab.Term.commitIfNoErrors? <| do - let tac ← `(by intros; try { arith }) - let v ← Lean.Elab.Term.elabTerm tac.raw (some vcondProofTy) - Lean.Elab.Term.synthesizeSyntheticMVarsNoPostponing - instantiateMVars v - - if let some e' := e then - -- The inferred variable condition. - let newCondition := mkAppNBeta e' args - let newCondition := mkAppNBeta newCondition (originalConstrVars.map (mkFVar ·.fvarId)) - - vcondData := vcondData.push (Sum.inr newCondition) - else - throwError "Variable condition {n} not found or inferred: \n {vcond} {constraints}." + vcondData := ← + withExistingLocalDecls originalConstrVars.toList do + let (e, _) ← Lean.Elab.Term.TermElabM.run <| Lean.Elab.Term.commitIfNoErrors? <| do + let tac ← `(by arith) + let v ← Lean.Elab.Term.elabTerm tac.raw (some vcond) + Lean.Elab.Term.synthesizeSyntheticMVarsNoPostponing + instantiateMVars v + + if let some e' := e then + return vcondData.push (Sum.inr e') + else + throwError "Variable condition {n} not found or inferred: \n {vcond} {constraints}." return (Tree.node vcondData childrenVCondData) | Tree.leaf _, Tree.leaf _ => pure (Tree.leaf ()) @@ -877,7 +871,7 @@ def mkProcessedAtomTree (objCurv : Curvature) (objFun : Expr) (constraints : Lis -- solution to solution but also the forward and backward maps. -- TODO: Better types for return type. def canonizeGoalFromSolutionExpr (goalExprs : Meta.SolutionExpr) : - MetaM (Expr × (Expr × Expr × Expr)) := do + MetaM (Expr × Expr × Expr) := do -- Extract objective and constraints from `goalExprs`. let (objFun, constraints, originalVarsDecls) ← withLambdaBody goalExprs.constraints fun p constraints => do @@ -971,25 +965,26 @@ def canonizeGoalFromSolutionExpr (goalExprs : Meta.SolutionExpr) : check res trace[Meta.debug] "second check passed" let res ← instantiateMVars res + check res + trace[Meta.debug] "instantiate mvars passed" return (forwardMap, backwardMap, res) - let newGoal ← mkFreshExprMVar none - - return (newGoal, (forwardMap, backwardMap, reduction)) + return (forwardMap, backwardMap, reduction) /-- -/ -def canonizeGoalFromExpr (goalExpr : Expr) : MetaM (Expr × (Expr × Expr × Expr)) := do +def canonizeGoalFromExpr (goalExpr : Expr) : MetaM (Expr × Expr × Expr) := do let goalExprs ← Meta.SolutionExpr.fromExpr goalExpr canonizeGoalFromSolutionExpr goalExprs /-- -/ -def canonizeGoal (goal : MVarId) : MetaM MVarId := do +def canonizeGoal (goal : MVarId) : MetaM (List MVarId) := do let goalExprs ← Meta.SolutionExpr.fromGoal goal - let (newGoal, (_forwardMap, _backwardMap, reduction)) ← canonizeGoalFromSolutionExpr goalExprs + let (_forwardMap, _backwardMap, reduction) ← canonizeGoalFromSolutionExpr goalExprs + let newGoal ← mkFreshExprMVar none let assignment := mkApp reduction newGoal check assignment goal.assign assignment - return newGoal.mvarId! + return [newGoal.mvarId!] end DCP @@ -1001,9 +996,7 @@ syntax (name := dcp) "dcp" : tactic @[tactic dcp] def evalDcp : Tactic := fun stx => match stx with -| `(tactic| dcp) => do - let goal ← Elab.Tactic.getMainGoal - replaceMainGoal [← DCP.canonizeGoal goal] +| `(tactic| dcp) => liftMetaTactic DCP.canonizeGoal | _ => throwUnsupportedSyntax end Tactic From ef3daca01e802106ce14bdd9b6a722131e98691d Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 16:56:29 -0500 Subject: [PATCH 161/189] chroe: problem goal not needed for `solve` --- CvxLean/Command/Solve.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/CvxLean/Command/Solve.lean b/CvxLean/Command/Solve.lean index 97f1adab..d840e56c 100644 --- a/CvxLean/Command/Solve.lean +++ b/CvxLean/Command/Solve.lean @@ -70,7 +70,7 @@ def getReducedProblemAndReduction (prob : Expr) trace[Meta.debug] "probOpt: {probOpt.objFun}" -- NOTE: We should get the value from this applied to the float sol point. - let (_, (forwardMap, backwardMap, probReduction)) ← DCP.canonizeGoalFromExpr probSol + let (forwardMap, backwardMap, probReduction) ← DCP.canonizeGoalFromExpr probSol let probReducedSol := match (← inferType probReduction) with | Expr.forallE _ r _ _ => r From 93ec0fb5d9ea4dbd66641a62263935756c16ae48 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 16:58:35 -0500 Subject: [PATCH 162/189] chore: disable `trace.Meta.debug`s --- CvxLean/Examples/CovarianceEstimation.lean | 2 -- CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean | 1 - CvxLean/Test/Convexify/DQCP.lean | 1 - CvxLean/Test/DCP/Dcp.lean | 2 -- CvxLean/Test/Other/Coeffs.lean | 2 -- CvxLean/Test/Problems/LogDet.lean | 10 ++++------ CvxLean/Test/Problems/SO.lean | 2 -- 7 files changed, 4 insertions(+), 16 deletions(-) diff --git a/CvxLean/Examples/CovarianceEstimation.lean b/CvxLean/Examples/CovarianceEstimation.lean index b0b3e927..618c23e9 100644 --- a/CvxLean/Examples/CovarianceEstimation.lean +++ b/CvxLean/Examples/CovarianceEstimation.lean @@ -58,8 +58,6 @@ reduction reduction₁₂/problem₂ (n : ℕ) (N : ℕ) (α : ℝ) (y : Fin N intro hR rw [nonsing_inv_nonsing_inv R hR.isUnit_det] -set_option trace.Meta.debug true - #print problem₂ set_option maxHeartbeats 20000000 diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean index 2ef9ff43..91ddf5b9 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean @@ -14,7 +14,6 @@ namespace CvxLean open Real -set_option trace.Meta.debug true in declare_atom huber [convex] (x : ℝ)? : huber x := vconditions implementationVars (v : ℝ) (w : ℝ) diff --git a/CvxLean/Test/Convexify/DQCP.lean b/CvxLean/Test/Convexify/DQCP.lean index 87a1f71c..0f2aa875 100644 --- a/CvxLean/Test/Convexify/DQCP.lean +++ b/CvxLean/Test/Convexify/DQCP.lean @@ -51,7 +51,6 @@ def hypersonicShapeDesign (a b : ℝ) := h2 : Δx ≤ 1 h3 : a * (1 / Δx) - (1 - b) * sqrt (1 - Δx ^ 2) ≤ 0 -set_option trace.Meta.debug true in time_cmd equivalence redqcp2/dqcp2 : hypersonicShapeDesign 0.05 0.65 := by unfold hypersonicShapeDesign convexify diff --git a/CvxLean/Test/DCP/Dcp.lean b/CvxLean/Test/DCP/Dcp.lean index a283a2e3..df74a178 100644 --- a/CvxLean/Test/DCP/Dcp.lean +++ b/CvxLean/Test/DCP/Dcp.lean @@ -16,7 +16,6 @@ noncomputable def testVCondInference : Solution <| dcp sorry -set_option trace.Meta.debug true noncomputable def test004 : Solution $ optimization (v w : ℝ) minimize exp v @@ -28,7 +27,6 @@ noncomputable def test004 : Solution $ dcp sorry -set_option trace.Meta.debug true noncomputable def test000 : Solution $ optimization (x : Fin 3 → ℝ) (y : ℝ) minimize y diff --git a/CvxLean/Test/Other/Coeffs.lean b/CvxLean/Test/Other/Coeffs.lean index d4299f82..90e9285a 100644 --- a/CvxLean/Test/Other/Coeffs.lean +++ b/CvxLean/Test/Other/Coeffs.lean @@ -6,8 +6,6 @@ section Coeffs open CvxLean Minimization -set_option trace.Meta.debug true - noncomputable def testVecExp1 : Solution $ optimization (x y : Fin 1 → ℝ) minimize (0 : ℝ) diff --git a/CvxLean/Test/Problems/LogDet.lean b/CvxLean/Test/Problems/LogDet.lean index 4e1d3dec..312b503d 100644 --- a/CvxLean/Test/Problems/LogDet.lean +++ b/CvxLean/Test/Problems/LogDet.lean @@ -5,9 +5,9 @@ section LogDet open CvxLean Minimization Real noncomputable def logDet1 := - optimization (X : Matrix (Fin 2) (Fin 2) ℝ) + optimization (X : Matrix (Fin 2) (Fin 2) ℝ) minimize (Matrix.trace X) - subject to + subject to h₁ : 10 ≤ log X.det h₂ : X.PosDef @@ -23,15 +23,13 @@ solve logDet1 #eval logDet1.solution 1 1 -- 148.413156 noncomputable def logDet2 := - optimization (X : Matrix (Fin 2) (Fin 2) ℝ) + optimization (X : Matrix (Fin 2) (Fin 2) ℝ) maximize (log X.det) - subject to + subject to h₁ : X.PosDef h₂ : X 0 0 + X 0 1 + X 1 1 ≤ 50 h₃ : X 0 1 = X 1 0 -set_option trace.Meta.debug true - solve logDet2 #print logDet2.reduced diff --git a/CvxLean/Test/Problems/SO.lean b/CvxLean/Test/Problems/SO.lean index aabfc049..9fb91cf6 100644 --- a/CvxLean/Test/Problems/SO.lean +++ b/CvxLean/Test/Problems/SO.lean @@ -27,8 +27,6 @@ def so2 := hx : 1 / 1000 ≤ x h : exp (-x) ≤ sqrt x -set_option trace.Meta.debug true - solve so2 #print so2.reduced From 4efe5821ae1ea3f301b1aa61f896a86fe53c6185 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 17:51:09 -0500 Subject: [PATCH 163/189] =?UTF-8?q?fix:=20`logSumExp=E2=82=82`=20in=20tran?= =?UTF-8?q?slation?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- CvxLean/Tactic/Convexify/Egg/FromMinimization.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean index 9814685b..eb9bad26 100644 --- a/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean +++ b/CvxLean/Tactic/Convexify/Egg/FromMinimization.lean @@ -78,7 +78,7 @@ def _root_.EggTree.opMap : HashMap String (String × Array EggTree.OpArgTag) := ("div", ("div", #[.arg, .arg])), ("quadOverLin", ("qol", #[.arg, .arg])), ("geoMean", ("geo", #[.arg, .arg])), - ("logSumExp", ("lse", #[.arg, .arg])), + ("logSumExp₂", ("lse", #[.arg, .arg])), ("norm2₂", ("norm2", #[.arg, .arg])) ] From 280b4e7a2d15a08decd67edf4f8dca52fb6d193d Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 17:51:22 -0500 Subject: [PATCH 164/189] feat: more robust `simp_or_rw` --- CvxLean/Tactic/Convexify/RewriteMapLibrary.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index ee5b775e..a5539d76 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -12,9 +12,9 @@ macro_rules match rule.raw[1] with | `(term| $e:term) => if symm then - `(tactic| (first | simp only [←$e:term] | repeat' (rw [←$e:term]))) + `(tactic| (try { simp only [←$e:term] } <;> repeat' (rw [←$e:term]))) else - `(tactic| (first | simp only [$e:term] | repeat' (rw [$e:term]))) + `(tactic| (try { simp only [$e:term] } <;> repeat' (rw [$e:term]))) namespace CvxLean From 93bd6d904daf50a5acc1af1837c95faf35a38469 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 18:03:12 -0500 Subject: [PATCH 165/189] test: formalize DCP quiz tests --- CvxLean/Test/Convexify/Quiz.lean | 106 +++++++++++++++++++++++++++++++ CvxLean/Test/Convexify/quiz.py | 85 +++++++++++++++++++++++++ 2 files changed, 191 insertions(+) create mode 100644 CvxLean/Test/Convexify/Quiz.lean create mode 100644 CvxLean/Test/Convexify/quiz.py diff --git a/CvxLean/Test/Convexify/Quiz.lean b/CvxLean/Test/Convexify/Quiz.lean new file mode 100644 index 00000000..a112669f --- /dev/null +++ b/CvxLean/Test/Convexify/Quiz.lean @@ -0,0 +1,106 @@ +import CvxLean.Command.Equivalence +import CvxLean.Tactic.DCP.AtomLibrary.All +import CvxLean.Tactic.Convexify.Convexify +import CvxLean.Test.Util.TimeCmd + +namespace Quiz + +noncomputable section + +open CvxLean Minimization Real + +def quiz1 := + optimization (x : ℝ) + minimize (x⁻¹⁻¹) + subject to + h : 0 < x + +time_cmd equivalence eq1/dcpquiz1 : quiz1 := by + convexify + +def quiz2 := + optimization (x y : ℝ) + minimize (-(log (exp (log (x + 42)) + exp (log y)))) + subject to + h₁ : 0 < x + h₂ : 0 < y + +time_cmd equivalence eq2/dcpquiz2 : quiz2 := by + convexify + +def quiz3 := + optimization (x : ℝ) + minimize ((sqrt x) ^ 2) + subject to + h : 0 ≤ x + +time_cmd equivalence eq3/dcpquiz3 : quiz3 := by + convexify + +def quiz4 := + optimization (x : ℝ) + minimize (-(abs (sqrt (abs (x + 42))))) + subject to + h : 0 ≤ x + +time_cmd equivalence eq4/dcpquiz4 : quiz4 := by + convexify + +def quiz5 := + optimization (x : ℝ) + minimize (1 / exp (x + 42)) + subject to + h : 1000 ≤ x -- irrelevant + +time_cmd equivalence eq5/dcpquiz5 : quiz5 := by + convexify + +def quiz6 := + optimization (x : ℝ) + minimize (-(log ((364 * x) ^ 2))) + subject to + h : 0 ≤ x + +time_cmd equivalence eq6/dcpquiz6 : quiz6 := by + convexify + +def quiz7 := + optimization (x : ℝ) + minimize ((sqrt ((x + 2) * (1 / x))) ^ 2) + subject to + h : 0 < x + +time_cmd equivalence eq7/dcpquiz7 : quiz7 := by + convexify + +def quiz8 := + optimization (x : ℝ) + minimize (-(log (abs x))) + subject to + h : 0 ≤ x + +time_cmd equivalence eq8/dcpquiz8 : quiz8 := by + convexify + +def quiz9 := + optimization (x y : ℝ) + minimize (1 / (((x⁻¹) ^ 2) / (y⁻¹))) + subject to + h₁ : 0 < x + h₂ : 0 < y + +time_cmd equivalence eq9/dcpquiz9 : quiz9 := by + convexify + +def quiz10 := + optimization (x : ℝ) + minimize ((log (exp x)) ^ 2) + subject to + h : 0 ≤ x + +time_cmd equivalence eq10/dcpquiz10 : quiz10 := by + convexify + +end + +end Quiz diff --git a/CvxLean/Test/Convexify/quiz.py b/CvxLean/Test/Convexify/quiz.py new file mode 100644 index 00000000..3897e6e6 --- /dev/null +++ b/CvxLean/Test/Convexify/quiz.py @@ -0,0 +1,85 @@ +import cvxpy as cp + +def quiz1(): + x = cp.Variable(pos=True) + + quiz1 = cp.inv_pos(cp.inv_pos(x)) + + assert(not quiz1.is_dcp()) + +def quiz2(): + x = cp.Variable(pos=True) + y = cp.Variable(pos=True) + + quiz2 = cp.inv_pos(cp.inv_pos(x) + cp.inv_pos(y)) + + assert(not quiz2.is_dcp()) + +def quiz3(): + x = cp.Variable(pos=True) + + quiz3 = cp.power(cp.sqrt(x), 2) + + assert(not quiz3.is_dcp()) + +def quiz4(): + x = cp.Variable(pos=True) + + quiz4 = -cp.abs(cp.sqrt(cp.abs(x))) + + assert(not quiz4.is_dcp()) + +def quiz5(): + x = cp.Variable() + + quiz5 = 1 / cp.exp(x) + + assert(not quiz5.is_dcp()) + +def quiz6(): + x = cp.Variable(pos=True) + + quiz6 = -cp.log(cp.power(364 * x, 2)) + + assert(not quiz6.is_dcp()) + +def quiz7(): + x = cp.Variable(pos=True) + + quiz7 = cp.power(cp.geo_mean(cp.vstack([x + 2, 1 / x])), 2) + + assert(not quiz7.is_dcp()) + +def quiz8(): + x = cp.Variable(pos=True) + + quiz8 = -cp.log(cp.abs(x)) + + assert(not quiz8.is_dcp()) + +def quiz9(): + x = cp.Variable(pos=True) + y = cp.Variable(pos=True) + + quiz9 = 1 / cp.quad_over_lin(1 / x, 1 / y) + + assert(not quiz9.is_dcp()) + +def quiz10(): + x = cp.Variable() + + quiz10 = cp.power(cp.log(cp.exp(x)), 2) + + assert(not quiz10.is_dcp()) + +if __name__ == "__main__": + quiz1() + quiz2() + quiz3() + quiz4() + quiz5() + quiz6() + quiz7() + quiz8() + quiz9() + quiz10() From f4c68e4b09b17b64f0ebc77009b89e037d1b28c9 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 18:04:01 -0500 Subject: [PATCH 166/189] test: add quiz tests to test suite --- CvxLean/Test/Convexify/All.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/CvxLean/Test/Convexify/All.lean b/CvxLean/Test/Convexify/All.lean index 0fd3f819..99408280 100644 --- a/CvxLean/Test/Convexify/All.lean +++ b/CvxLean/Test/Convexify/All.lean @@ -1,5 +1,6 @@ -import CvxLean.Test.Convexify.MainExample import CvxLean.Test.Convexify.Unit -import CvxLean.Test.Convexify.DGP import CvxLean.Test.Convexify.AlmostDGP +import CvxLean.Test.Convexify.DGP import CvxLean.Test.Convexify.DQCP +import CvxLean.Test.Convexify.MainExample +import CvxLean.Test.Convexify.Quiz From d863d8a61872266ed435bd0ebf204bcb2c90936a Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Mon, 11 Dec 2023 18:37:23 -0500 Subject: [PATCH 167/189] chore: print term JSON --- CvxLean/Tactic/Convexify/Convexify.lean | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/CvxLean/Tactic/Convexify/Convexify.lean b/CvxLean/Tactic/Convexify/Convexify.lean index 76ad52ee..2c12ab5e 100644 --- a/CvxLean/Tactic/Convexify/Convexify.lean +++ b/CvxLean/Tactic/Convexify/Convexify.lean @@ -272,9 +272,10 @@ def evalConvexify : Tactic := fun stx => match stx with constr := gStr.constr.filter (fun (h, _) => !constrsToIgnore.contains h) } -- Prepare egg request. + let eggMinimization := EggMinimization.ofOCTree gStr let eggRequest : EggRequest := { domains := varDomainConstrs.data, - target := EggMinimization.ofOCTree gStr + target := eggMinimization } try @@ -287,6 +288,7 @@ def evalConvexify : Tactic := fun stx => match stx with dbg_trace s!"Number of steps: {steps.size}." let size := (gStr.map fun (_, t) => t.size).fold 0 Nat.add dbg_trace s!"Term size: {size}." + dbg_trace s!"Term JSON: {eggMinimization.toJson}." -- Apply steps. let mut g ← getMainGoal From f31e4a6d8a82bc27cbaf7fbf2607da035046e78e Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Tue, 12 Dec 2023 10:45:58 -0500 Subject: [PATCH 168/189] chore: print problem names for `convexify` evaluation --- CvxLean/Tactic/Convexify/Convexify.lean | 3 +++ CvxLean/Test/Convexify/DQCP.lean | 4 ---- 2 files changed, 3 insertions(+), 4 deletions(-) diff --git a/CvxLean/Tactic/Convexify/Convexify.lean b/CvxLean/Tactic/Convexify/Convexify.lean index 2c12ab5e..783537c6 100644 --- a/CvxLean/Tactic/Convexify/Convexify.lean +++ b/CvxLean/Tactic/Convexify/Convexify.lean @@ -238,9 +238,12 @@ def evalConvexify : Tactic := fun stx => match stx with -- again. if let Expr.const n _ := gExprRaw then unfoldTarget n + dbg_trace s!"Convexify problem name: {n}" normNumCleanUp (useSimp := false) let gTarget ← getMainTarget gExprRaw ← liftM <| Meta.getExprRawFromGoal isEquiv gTarget + else + dbg_trace s!"Convexify problem name: unknown" -- Get `MinmizationExpr`. let gExpr ← MinimizationExpr.fromExpr gExprRaw diff --git a/CvxLean/Test/Convexify/DQCP.lean b/CvxLean/Test/Convexify/DQCP.lean index 0f2aa875..63e97c87 100644 --- a/CvxLean/Test/Convexify/DQCP.lean +++ b/CvxLean/Test/Convexify/DQCP.lean @@ -55,10 +55,6 @@ time_cmd equivalence redqcp2/dqcp2 : hypersonicShapeDesign 0.05 0.65 := by unfold hypersonicShapeDesign convexify --- TODO: Error in solve (or dcp): --- Lean server printed an error: --- PANIC at Lean.Expr.mvarId! Lean.Expr:1067:14: mvar expected --- PANIC at Lean.MetavarContext.getDecl Lean.MetavarContext:398:17: unknown metavariable solve dqcp2 #print dqcp2.reduced From f94b6fcaad4b56449b39b09cb27cefc16b927473 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 13 Dec 2023 13:59:19 -0500 Subject: [PATCH 169/189] test: start checker tests --- egg-convexify/tests/test_domain.rs | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/egg-convexify/tests/test_domain.rs b/egg-convexify/tests/test_domain.rs index 488877b6..72993f15 100644 --- a/egg-convexify/tests/test_domain.rs +++ b/egg-convexify/tests/test_domain.rs @@ -686,3 +686,11 @@ fn exp_neg() { let expected = Domain::make_oo(domain::zero(), domain::one()); assert!(result.eq(&expected)); } + +/* Checkers tests. TODO: more. */ + +#[test] +fn is_nonneg_1() { + // [1, 1] is non-negative + assert!(domain::is_nonneg(&Domain::make_singleton(1.0))); +} From 545aa1f33ebe1578177a636b221672227981ffd9 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 13 Dec 2023 14:08:26 -0500 Subject: [PATCH 170/189] feat: more refined even powers --- egg-convexify/src/domain.rs | 22 ++++++++++++++++------ 1 file changed, 16 insertions(+), 6 deletions(-) diff --git a/egg-convexify/src/domain.rs b/egg-convexify/src/domain.rs index 58632797..d085b72f 100644 --- a/egg-convexify/src/domain.rs +++ b/egg-convexify/src/domain.rs @@ -79,7 +79,7 @@ impl Domain { } pub fn make_ci(lo: Float) -> Self { - Domain::make_from_endpoints(lo, inf(), false, false) + Domain::make_from_endpoints(lo, inf(), false, true) } pub fn make_oc(lo: Float, hi: Float) -> Self { @@ -91,19 +91,19 @@ impl Domain { } pub fn make_oi(lo: Float) -> Self { - Domain::make_from_endpoints(lo, inf(), true, false) + Domain::make_from_endpoints(lo, inf(), true, true) } pub fn make_ic(hi: Float) -> Self { - Domain::make_from_endpoints(neg_inf(), hi, false, false) + Domain::make_from_endpoints(neg_inf(), hi, true, false) } pub fn make_io(hi: Float) -> Self { - Domain::make_from_endpoints(neg_inf(), hi, false, true) + Domain::make_from_endpoints(neg_inf(), hi, true, true) } pub fn make_ii() -> Self { - Domain::make_from_endpoints(neg_inf(), inf(), false, false) + Domain::make_from_endpoints(neg_inf(), inf(), true, true) } /* Comparisons. */ @@ -848,7 +848,17 @@ pub fn pow(d1: &Domain, d2: &Domain) -> Domain { } else if n == 1 { return d1.clone(); } else if n % 2 == 0 { - return nonneg_dom(); + let d = abs(d1); + let lo = d.lo_float(); + let hi = d.hi_float(); + let c = &Float::with_val(F64_PREC, f); + let l = d.lo_open; + let r = d.hi_open; + let res = + perform_pow( + lo, c, hi, c, + l, false, r, false); + return res; } else { return free_dom(); } From ec577eedbb8a08287b608209cb703eeff6b6590d Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 13 Dec 2023 14:19:16 -0500 Subject: [PATCH 171/189] feat: `rpow_eq_mul_rpow_pred` --- CvxLean/Lib/Math/Data/Real.lean | 4 ++++ CvxLean/Tactic/Convexify/RewriteMapLibrary.lean | 3 +++ egg-convexify/src/rules.rs | 3 +++ 3 files changed, 10 insertions(+) diff --git a/CvxLean/Lib/Math/Data/Real.lean b/CvxLean/Lib/Math/Data/Real.lean index 7e60c83f..683d6d29 100644 --- a/CvxLean/Lib/Math/Data/Real.lean +++ b/CvxLean/Lib/Math/Data/Real.lean @@ -75,6 +75,10 @@ lemma binomial_two (x y : ℝ) : (x + y) ^ 2 = x ^ 2 + (2 * (x * y) + y ^ 2) := by simp only [rpow_two]; ring +lemma rpow_eq_mul_rpow_pred {x y : ℝ} (hx : x ≠ 0) : + x ^ y = x * (x ^ (y - 1)) := by + conv => left; rw [(by ring : y = (y - 1) + 1), rpow_add_one hx, mul_comm] + lemma exp_neg_eq_one_div (x : ℝ) : exp (-x) = 1 / exp x := by rw [exp_neg, inv_eq_one_div] diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index a5539d76..74bf5aa2 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -241,6 +241,9 @@ register_rewrite_map "pow_half_two-rev"; "?a" => "(pow (pow ?a 0.5) 2)" := register_rewrite_map "binomial_two"; "(pow (add ?a ?b) 2)" => "(add (pow ?a 2) (add (mul 2 (mul ?a ?b)) (pow ?b 2)))" := simp_or_rw [Real.binomial_two]; +register_rewrite_map "rpow_eq_mul_rpow_pred"; "(pow ?a ?b)" => "(mul ?a (pow ?a (sub ?b 1)))" := + simp_or_rw [Real.rpow_eq_mul_rpow_pred (by positivity!)]; + register_rewrite_map "inv_eq_pow_neg_one"; "(inv ?a)" => "(pow ?a (neg 1))" := simp_or_rw [Real.inv_eq_pow_neg_one (by positivity!)]; diff --git a/egg-convexify/src/rules.rs b/egg-convexify/src/rules.rs index 379fb222..04c0e5d2 100644 --- a/egg-convexify/src/rules.rs +++ b/egg-convexify/src/rules.rs @@ -191,6 +191,9 @@ pub fn rules() -> Vec> { vec![ rw!("binomial_two"; "(pow (add ?a ?b) 2)" => "(add (pow ?a 2) (add (mul 2 (mul ?a ?b)) (pow ?b 2)))"), + rw!("rpow_eq_mul_rpow_pred"; "(pow ?a ?b)" => "(mul ?a (pow ?a (sub ?b 1)))" + if is_not_zero("?a")), + rw!("inv_eq_pow_neg_one"; "(inv ?a)" => "(pow ?a (neg 1))" if is_not_zero("?a")), From 580f866192ab0d8a861c1d64a0d5576185bb6762 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 13 Dec 2023 15:25:14 -0500 Subject: [PATCH 172/189] wip: Stanford tests --- egg-convexify/tests/test_stanford.rs | 38 ++++++++++++++++++---------- 1 file changed, 24 insertions(+), 14 deletions(-) diff --git a/egg-convexify/tests/test_stanford.rs b/egg-convexify/tests/test_stanford.rs index 7ea6f753..2f1f0f82 100644 --- a/egg-convexify/tests/test_stanford.rs +++ b/egg-convexify/tests/test_stanford.rs @@ -7,6 +7,7 @@ use egg_convexify::domain; use egg_convexify::test_util::{*}; +// TODO: Failing because of qol curvature check simplification. #[test] fn test_3_32() { // 1 / (x * y) = (x^-0.5)^2 / y @@ -16,6 +17,7 @@ fn test_3_32() { "(div 1 (mul (var x) (var y)))"); } +// TODO: Failing because of norm2 curvature check simplification. #[test] fn test_3_33() { // sqrt(1 + x^4 / y) = sqrt(1^2 + (x^2 / y)^2) @@ -28,6 +30,7 @@ fn test_3_33() { ); } +// TODO: Failing because of norm2 curvature check simplification. #[test] fn test_3_36_a() { // sqrt(1 + 4x^2 + 16y^2) = sqrt((2x)^2 + (sqrt(1^2 + (4y)^2))^2) @@ -37,13 +40,13 @@ fn test_3_36_a() { "(sqrt (add 1 (add (mul 4 (pow (var x) 2)) (mul 16 (pow (var y) 2)))))"); } -// #[test] -// fn test_3_36_c() { -// // log(e^(2x + 3) + e^(4y + 5)) = lse(2x + 3, 4y + 5) -// convexify_check_expression_with_domain( -// vec![("x", domain::free_dom()), ("y", domain::free_dom())], -// "(log (add (exp (add (mul 2 (var x)) 3)) (exp (add (mul 4 (var y)) 5))))"); -// } +#[test] +fn test_3_36_c() { + // log(e^(2x + 3) + e^(4y + 5)) = lse(2x + 3, 4y + 5) + convexify_check_expression_with_domain( + vec![("x", domain::free_dom()), ("y", domain::free_dom())], + "(log (add (exp (add (mul 2 (var x)) 3)) (exp (add (mul 4 (var y)) 5))))"); +} #[test] fn test_3_38_e() { @@ -54,11 +57,7 @@ fn test_3_38_e() { "(neg (pow (add (sqrt (var x)) (sqrt (var y))) 2))"); } -#[test] -fn test_3_67() { - // Generalizaiton of 3.28. Works for n = 3, times out for n >= 4. - // (sqrt(x_1) + ... + sqrt(x_n))^2 - // ... = sum_{i <= n} x_i + 2 * sum_{i < j <= n} geo(x_i, x_j) +fn test_3_67_aux(n: usize) { let build_domain = |n| { if n < 2 { panic!("n must be >= 2"); @@ -86,11 +85,22 @@ fn test_3_67() { } format!("(neg (pow {} 2))", t) }; - let domain_pre = build_domain(3).clone(); + let domain_pre = build_domain(n).clone(); let domain = domain_pre .iter() .map(|(s,d)| (s.as_str(), d.clone())) .collect::>(); - convexify_check_expression_with_domain(domain, &build_term(3)); + convexify_check_expression_with_domain(domain, &build_term(n)); +} + +#[test] +fn test_3_67() { + // Generalizaiton of 3.28. Works for n = 3, times out for n >= 4. + // (sqrt(x_1) + ... + sqrt(x_n))^2 + // ... = sum_{i <= n} x_i + 2 * sum_{i < j <= n} geo(x_i, x_j) + test_3_67_aux(3); + // test_3_67_aux(4); + // test_3_67_aux(5); + // test_3_67_aux(6); } From 4946b02a2f39767b073f4e0ffc840a93a185b870 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 13 Dec 2023 15:58:11 -0500 Subject: [PATCH 173/189] test: benchmarking 3.67 --- egg-convexify/tests/test_stanford.rs | 27 ++++++++++++++++----------- 1 file changed, 16 insertions(+), 11 deletions(-) diff --git a/egg-convexify/tests/test_stanford.rs b/egg-convexify/tests/test_stanford.rs index 2f1f0f82..ede4a56e 100644 --- a/egg-convexify/tests/test_stanford.rs +++ b/egg-convexify/tests/test_stanford.rs @@ -3,8 +3,9 @@ Tests from the additional exercises in the Stanford Convex Optimization course: https://github.com/cvxgrp/cvxbook_additional_exercises !*/ -use egg_convexify::domain; +use std::time::Instant; +use egg_convexify::domain; use egg_convexify::test_util::{*}; // TODO: Failing because of qol curvature check simplification. @@ -57,7 +58,10 @@ fn test_3_38_e() { "(neg (pow (add (sqrt (var x)) (sqrt (var y))) 2))"); } -fn test_3_67_aux(n: usize) { +fn test_3_67_aux(n: usize, node_limit: usize) { + // Generalizaiton of 3.28. Works for n = 3,4,5,6,7 + // (sqrt(x_1) + ... + sqrt(x_n))^2 + // ... = sum_{i <= n} x_i + 2 * sum_{i < j <= n} geo(x_i, x_j) let build_domain = |n| { if n < 2 { panic!("n must be >= 2"); @@ -91,16 +95,17 @@ fn test_3_67_aux(n: usize) { .iter() .map(|(s,d)| (s.as_str(), d.clone())) .collect::>(); - convexify_check_expression_with_domain(domain, &build_term(n)); + convexify_check_expression_with_domain_and_node_limit(domain, &build_term(n), node_limit); } #[test] fn test_3_67() { - // Generalizaiton of 3.28. Works for n = 3, times out for n >= 4. - // (sqrt(x_1) + ... + sqrt(x_n))^2 - // ... = sum_{i <= n} x_i + 2 * sum_{i < j <= n} geo(x_i, x_j) - test_3_67_aux(3); - // test_3_67_aux(4); - // test_3_67_aux(5); - // test_3_67_aux(6); -} + for n in 2..10 { + let now = Instant::now(); + { + test_3_67_aux(n, 20000 * n); + } + let elapsed = now.elapsed(); + println!("Time for 3.67 (n={}): {:.2?}", n, elapsed); + } +} \ No newline at end of file From 0fd76c9e8ab6a3269122e891cb2319aa2c702e28 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 13 Dec 2023 15:58:24 -0500 Subject: [PATCH 174/189] feat: extend test util --- egg-convexify/src/test_util.rs | 25 ++++++++++++++++++++++--- 1 file changed, 22 insertions(+), 3 deletions(-) diff --git a/egg-convexify/src/test_util.rs b/egg-convexify/src/test_util.rs index c4d4f6ec..2559151d 100644 --- a/egg-convexify/src/test_util.rs +++ b/egg-convexify/src/test_util.rs @@ -4,7 +4,7 @@ use domain::Domain as Domain; use crate::extract; use extract::Minimization as Minimization; use extract::get_steps as get_steps; -use extract::get_steps_from_string as get_steps_from_string; +use extract::get_steps_from_string_maybe_node_limit as get_steps_from_string_maybe_node_limit; fn make(obj: &str, constrs: Vec<&str>) -> Minimization { let mut constrs_s = Vec::new(); @@ -49,10 +49,13 @@ pub fn convexify_check_and_print(obj: &str, constrs: Vec<&str>) { // Used to test out-of-context expressions. -fn convexify_check_expression_with_domain_maybe_print(domains : Vec<(&str, Domain)>, s: &str, print: bool) { +fn convexify_check_expression_with_domain_maybe_print_maybe_node_limit( + domains : Vec<(&str, Domain)>, + s: &str, print: bool, + node_limit: Option) { let domains = domains.iter().map(|(s, d)| ((*s).to_string(), d.clone())).collect(); - let steps = get_steps_from_string(s, domains, true); + let steps = get_steps_from_string_maybe_node_limit(s, domains, true, node_limit); if steps.is_none() { panic!("Test failed, could not rewrite target into DCP form."); } @@ -61,6 +64,10 @@ fn convexify_check_expression_with_domain_maybe_print(domains : Vec<(&str, Domai } } +fn convexify_check_expression_with_domain_maybe_print(domains : Vec<(&str, Domain)>, s: &str, print: bool) { + convexify_check_expression_with_domain_maybe_print_maybe_node_limit(domains, s, print, None); +} + pub fn convexify_check_expression_with_domain(domains : Vec<(&str, Domain)>,s: &str) { convexify_check_expression_with_domain_maybe_print(domains, s, false); } @@ -69,6 +76,18 @@ pub fn convexify_check_expression_with_domain_and_print(domains : Vec<(&str, Dom convexify_check_expression_with_domain_maybe_print(domains, s, true); } +fn convexify_check_expression_with_domain_and_node_limit_maybe_print(domains : Vec<(&str, Domain)>, s: &str, print: bool, node_limit: usize) { + convexify_check_expression_with_domain_maybe_print_maybe_node_limit(domains, s, print, Some(node_limit)); +} + +pub fn convexify_check_expression_with_domain_and_node_limit(domains : Vec<(&str, Domain)>,s: &str, node_limit: usize) { + convexify_check_expression_with_domain_and_node_limit_maybe_print(domains, s, false, node_limit); +} + +pub fn convexify_check_expression_with_domain_and_node_limit_and_print(domains : Vec<(&str, Domain)>,s: &str, node_limit: usize) { + convexify_check_expression_with_domain_and_node_limit_maybe_print(domains, s, true, node_limit); +} + pub fn convexify_check_expression(s: &str) { convexify_check_expression_with_domain_maybe_print(vec![], s, false); } From d8147bb6d9016f9f4a8385714c013977ad3a7380 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 13 Dec 2023 15:58:35 -0500 Subject: [PATCH 175/189] feat: allow user-specified node limits --- egg-convexify/src/extract.rs | 17 ++++++++++++++++- 1 file changed, 16 insertions(+), 1 deletion(-) diff --git a/egg-convexify/src/extract.rs b/egg-convexify/src/extract.rs index fffd5f2b..5306d4a3 100644 --- a/egg-convexify/src/extract.rs +++ b/egg-convexify/src/extract.rs @@ -149,6 +149,14 @@ pub fn get_steps(prob: Minimization, domains_vec: Vec<(String, Domain)>, debug: } pub fn get_steps_from_string(prob_s: &str, domains_vec: Vec<(String, Domain)>, debug: bool) -> Option> { + get_steps_from_string_maybe_node_limit(prob_s, domains_vec, debug, None) +} + +pub fn get_steps_from_string_maybe_node_limit( + prob_s: &str, + domains_vec: Vec<(String, Domain)>, + debug: bool, + node_limit: Option) -> Option> { let expr: RecExpr = prob_s.parse().unwrap(); // Process domains, intersecting domains assigned to the same variable. @@ -172,7 +180,14 @@ pub fn get_steps_from_string(prob_s: &str, domains_vec: Vec<(String, Domain)>, d } } - for node_limit in [2500, 5000, 10000, 25000, 50000] { + // Specify a node limit, or try with default ones. + let node_limits = + if let Some(n) = node_limit { + vec![n] + } else { + vec![2500, 5000, 10000, 25000, 50000] + }; + for node_limit in node_limits { let analysis = Meta { domains : domains.clone() }; From 26254483d2d87af85d6b059530a6e96e43d746c6 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 13 Dec 2023 16:02:42 -0500 Subject: [PATCH 176/189] feat: `sqrt_mul` --- CvxLean/Tactic/Convexify/RewriteMapLibrary.lean | 6 ++++++ egg-convexify/src/rules.rs | 6 ++++++ 2 files changed, 12 insertions(+) diff --git a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean index 74bf5aa2..141045bc 100644 --- a/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean +++ b/CvxLean/Tactic/Convexify/RewriteMapLibrary.lean @@ -225,6 +225,12 @@ register_rewrite_map "sqrt_eq_rpow" ; "(sqrt ?a)" => "(pow ?a 0.5)" := register_rewrite_map "sqrt_eq_rpow-rev" ; "(pow ?a 0.5)" => "(sqrt ?a)" := simp_or_rw [←Real.sqrt_eq_rpow]; +register_rewrite_map "sqrt_mul" ; "(sqrt (mul ?a ?b))" => "(mul (sqrt ?a) (sqrt ?b))" := + simp_or_rw [Real.sqrt_mul (by positivity!)]; + +register_rewrite_map "sqrt_mul-rev" ; "(mul (sqrt ?a) (sqrt ?b))" => "(sqrt (mul ?a ?b))" := + simp_or_rw [←Real.sqrt_mul (by positivity!)]; + register_rewrite_map "pow_two"; "(pow ?a 2)" => "(mul ?a ?a)" := (norm_cast; simp_or_rw [pow_two]); diff --git a/egg-convexify/src/rules.rs b/egg-convexify/src/rules.rs index 04c0e5d2..549581be 100644 --- a/egg-convexify/src/rules.rs +++ b/egg-convexify/src/rules.rs @@ -181,6 +181,12 @@ pub fn rules() -> Vec> { vec![ rw!("sqrt_eq_rpow-rev"; "(pow ?a 0.5)" => "(sqrt ?a)"), + rw!("sqrt_mul"; "(sqrt (mul ?a ?b))" => "(mul (sqrt ?a) (sqrt ?b))" + if is_ge_zero("?a") if is_ge_zero("?b")), + + rw!("sqrt_mul-rev"; "(mul (sqrt ?a) (sqrt ?b))" => "(sqrt (mul ?a ?b))" + if is_ge_zero("?a") if is_ge_zero("?b")), + rw!("pow_two"; "(pow ?a 2)" => "(mul ?a ?a)"), rw!("pow_two-rev"; "(mul ?a ?a)" => "(pow ?a 2)"), From 44a8d94340382ce498eeb130a45533acc6eba363 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 13 Dec 2023 16:37:34 -0500 Subject: [PATCH 177/189] test: some statistics for 3.67 --- egg-convexify/tests/test_stanford.rs | 12 +++++++++++- 1 file changed, 11 insertions(+), 1 deletion(-) diff --git a/egg-convexify/tests/test_stanford.rs b/egg-convexify/tests/test_stanford.rs index ede4a56e..60e6b25b 100644 --- a/egg-convexify/tests/test_stanford.rs +++ b/egg-convexify/tests/test_stanford.rs @@ -100,7 +100,7 @@ fn test_3_67_aux(n: usize, node_limit: usize) { #[test] fn test_3_67() { - for n in 2..10 { + for n in 2..11 { let now = Instant::now(); { test_3_67_aux(n, 20000 * n); @@ -108,4 +108,14 @@ fn test_3_67() { let elapsed = now.elapsed(); println!("Time for 3.67 (n={}): {:.2?}", n, elapsed); } + // node_limit = 20000 * n + // n=2: 32359 nodes, 8 steps, 7.82s + // n=3: 52870 nodes, 28 steps, 23.97s + // n=4: 63224 nodes, 37 steps, 31.05s + // n=5: 102470 nodes, 87 steps, 46.83s + // n=6: 106484 nodes, 132 steps, 129.44s + // n=7: 135393 nodes, 282 steps, 152.53s + // n=8: 157549 nodes, 566 steps, 269.90s + // n=9: 174177 nodes, 511 steps, 307.08s + // n=10: 200372 nodes, 716 steps, 173.13s } \ No newline at end of file From d690d4552defc732133e09d713185dc5c4ffef2c Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Wed, 13 Dec 2023 16:38:45 -0500 Subject: [PATCH 178/189] chore: ignore plots --- .gitignore | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/.gitignore b/.gitignore index d15ee674..a7770089 100644 --- a/.gitignore +++ b/.gitignore @@ -1,7 +1,6 @@ .DS_Store /build -/evaluation /solver/*.cbf /solver/*.sol /egg-convexify/*.dot @@ -12,4 +11,7 @@ *.olean /playground +/evaluation +/plots + /CvxLean/Tactic/Solver/Mosek/Path.lean From fc8c080c2b5b14db7ea70a12147c2f9e672984ce Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 15 Dec 2023 11:13:33 -0500 Subject: [PATCH 179/189] fix: issue with empty `solver` directory --- CvxLean/Tactic/Solver/Conic.lean | 3 ++- CvxLean/Test/Problems/SO.lean | 2 +- 2 files changed, 3 insertions(+), 2 deletions(-) diff --git a/CvxLean/Tactic/Solver/Conic.lean b/CvxLean/Tactic/Solver/Conic.lean index 54cb94fa..c5c1aede 100644 --- a/CvxLean/Tactic/Solver/Conic.lean +++ b/CvxLean/Tactic/Solver/Conic.lean @@ -113,7 +113,8 @@ unsafe def conicSolverFromValues (goalExprs : SolutionExpr) -- TODO: locking? let outputPath := "solver/problem.sol" - IO.FS.withFile outputPath IO.FS.Mode.read fun handle => do + IO.FS.writeFile outputPath "" + IO.FS.withFile outputPath IO.FS.Mode.readWrite fun handle => do -- Run solver. let out ← IO.Process.output { cmd := "mosek", diff --git a/CvxLean/Test/Problems/SO.lean b/CvxLean/Test/Problems/SO.lean index 9fb91cf6..9c74b3aa 100644 --- a/CvxLean/Test/Problems/SO.lean +++ b/CvxLean/Test/Problems/SO.lean @@ -10,7 +10,7 @@ noncomputable def so1 := subject to c1 : y = 2 * x - 3 c2 : x ^ 2 ≤ 2 - c3 : 0 ≤ x - y -- TODO: error if 1 / 1000 ≤ x - y + c3 : 0 ≤ x - y solve so1 From 0be8102dd1be3d56092ab5ed367e48ae907b7e62 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 15 Dec 2023 11:43:43 -0500 Subject: [PATCH 180/189] fix: create MOSEK files in `build.sh` --- CvxLean/Tactic/Solver/Conic.lean | 21 ++++++++++----------- build.sh | 2 ++ 2 files changed, 12 insertions(+), 11 deletions(-) diff --git a/CvxLean/Tactic/Solver/Conic.lean b/CvxLean/Tactic/Solver/Conic.lean index c5c1aede..21d37417 100644 --- a/CvxLean/Tactic/Solver/Conic.lean +++ b/CvxLean/Tactic/Solver/Conic.lean @@ -101,20 +101,19 @@ unsafe def conicSolverFromValues (goalExprs : SolutionExpr) cbf := cbf.setScalarConstraints (CBF.ConeProduct.mk n groupedCones.length groupedCones) - -- Write input. - let inputPath := "solver/problem.cbf" - IO.FS.writeFile inputPath (ToString.toString cbf) - - -- Adjust path to MOSEK. - let p := if let some p' := ← IO.getEnv "PATH" then - if mosekBinPath != "" then p' ++ ":" ++ mosekBinPath else p' - else - mosekBinPath - -- TODO: locking? let outputPath := "solver/problem.sol" - IO.FS.writeFile outputPath "" IO.FS.withFile outputPath IO.FS.Mode.readWrite fun handle => do + -- Write input. + let inputPath := "solver/problem.cbf" + IO.FS.writeFile inputPath (ToString.toString cbf) + + -- Adjust path to MOSEK. + let p := if let some p' := ← IO.getEnv "PATH" then + if mosekBinPath != "" then p' ++ ":" ++ mosekBinPath else p' + else + mosekBinPath + -- Run solver. let out ← IO.Process.output { cmd := "mosek", diff --git a/build.sh b/build.sh index 20db3b5b..8f8e6c7f 100755 --- a/build.sh +++ b/build.sh @@ -1,4 +1,6 @@ #!/bin/bash +touch solver/problem.cbf +touch solver/problem.sol lake update lake exe cache get lake build EggConvexify From 6b70358eff7bd0ddfc90e79039017ef6f2049794 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 15 Dec 2023 12:31:31 -0500 Subject: [PATCH 181/189] test: add `gp8` to egg tests --- egg-convexify/tests/{test_gp.rs => test_dgp.rs} | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) rename egg-convexify/tests/{test_gp.rs => test_dgp.rs} (66%) diff --git a/egg-convexify/tests/test_gp.rs b/egg-convexify/tests/test_dgp.rs similarity index 66% rename from egg-convexify/tests/test_gp.rs rename to egg-convexify/tests/test_dgp.rs index bbc8979b..fd84be86 100644 --- a/egg-convexify/tests/test_gp.rs +++ b/egg-convexify/tests/test_dgp.rs @@ -46,3 +46,19 @@ fn test_gp6() { "(le (div (exp (var d)) (exp (var w))) (param δ))" ]); } + +#[test] +fn test_gp8() { + convexify_check( + "(mul (mul 2 (exp (var A))) (norm2 (exp (var w)) (exp (var h))))", + vec![ + "(le (mul (div (mul 10 (norm2 (exp (var w)) (exp (var h)))) 2) (exp (var h))) (mul (div 1 2) (exp (var A))))", + "(le (mul (div (mul 20 (norm2 (exp (var w)) (exp (var h)))) 2) (exp (var w))) (mul (div 1 2) (exp (var A))))", + "(le 1 (exp (var h)))", + "(le (exp (var h)) 100)", + "(le 1 (exp (var w)))", + "(le (exp (var w)) 100)", + "(le (mul (div 11 10) (exp (var r))) (sqrt (add (div (exp (var A)) (div 314159 50000)) (pow (exp (var r)) 2))))", + "(le (sqrt (add (div (exp (var A)) (div 314159 50000)) (pow (exp (var r)) 2))) 10)" + ]); +} From 4682a76e684ef258907075aa22eed0ee955aee9e Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 15 Dec 2023 12:32:08 -0500 Subject: [PATCH 182/189] chore: clean `egg` directory in `build.sh` --- build.sh | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/build.sh b/build.sh index 8f8e6c7f..79afd7b0 100755 --- a/build.sh +++ b/build.sh @@ -3,6 +3,7 @@ touch solver/problem.cbf touch solver/problem.sol lake update lake exe cache get +lake run EggClean lake build EggConvexify lake build CvxLean -lake build CvxLeanTest \ No newline at end of file +lake build CvxLeanTest From eef8b27cd76241f9abf070231da686ff09f050d7 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 15 Dec 2023 12:32:22 -0500 Subject: [PATCH 183/189] fix: remove power rule --- egg-convexify/src/rules.rs | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/egg-convexify/src/rules.rs b/egg-convexify/src/rules.rs index 549581be..6fdf2cfc 100644 --- a/egg-convexify/src/rules.rs +++ b/egg-convexify/src/rules.rs @@ -197,8 +197,8 @@ pub fn rules() -> Vec> { vec![ rw!("binomial_two"; "(pow (add ?a ?b) 2)" => "(add (pow ?a 2) (add (mul 2 (mul ?a ?b)) (pow ?b 2)))"), - rw!("rpow_eq_mul_rpow_pred"; "(pow ?a ?b)" => "(mul ?a (pow ?a (sub ?b 1)))" - if is_not_zero("?a")), + // rw!("rpow_eq_mul_rpow_pred"; "(pow ?a ?b)" => "(mul ?a (pow ?a (sub ?b 1)))" + // if is_not_zero("?a")), rw!("inv_eq_pow_neg_one"; "(inv ?a)" => "(pow ?a (neg 1))" if is_not_zero("?a")), From 00e19e0aaf118e97da2f463b171530a7edb5a1f5 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 15 Dec 2023 12:32:30 -0500 Subject: [PATCH 184/189] chore: increase node limit --- egg-convexify/src/extract.rs | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/egg-convexify/src/extract.rs b/egg-convexify/src/extract.rs index 5306d4a3..1641b35c 100644 --- a/egg-convexify/src/extract.rs +++ b/egg-convexify/src/extract.rs @@ -185,7 +185,7 @@ pub fn get_steps_from_string_maybe_node_limit( if let Some(n) = node_limit { vec![n] } else { - vec![2500, 5000, 10000, 25000, 50000] + vec![2500, 5000, 10000, 20000, 50000, 80000] }; for node_limit in node_limits { let analysis = Meta { From 700c909da2a16bc51a547e5459899a6560f9d2b0 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 15 Dec 2023 13:17:39 -0500 Subject: [PATCH 185/189] chore: bump to v4.4.0-rc1 --- .github/workflows/main.yml | 2 +- .gitignore | 4 +- lake-manifest.json | 117 ++++++++++++++++++------------------- lakefile.lean | 3 +- lean-toolchain | 2 +- 5 files changed, 63 insertions(+), 65 deletions(-) diff --git a/.github/workflows/main.yml b/.github/workflows/main.yml index 19c05d8d..c5574d58 100644 --- a/.github/workflows/main.yml +++ b/.github/workflows/main.yml @@ -27,7 +27,7 @@ jobs: run: | set -o pipefail curl -sSfL https://github.com/leanprover/elan/releases/download/v1.4.5/elan-x86_64-unknown-linux-gnu.tar.gz | tar xz - ./elan-init -y --default-toolchain leanprover/lean4:v4.3.0-rc1 + ./elan-init -y --default-toolchain leanprover/lean4:v4.4.0-rc1 echo "$HOME/.elan/bin" >> $GITHUB_PATH - name: Install Rust and Cargo run: | diff --git a/.gitignore b/.gitignore index a7770089..bdbc7386 100644 --- a/.gitignore +++ b/.gitignore @@ -1,13 +1,11 @@ .DS_Store +/.lake /build /solver/*.cbf /solver/*.sol /egg-convexify/*.dot /egg-convexify/*.svg -/ffi/build -/ffi/lake-packages -/lake-packages *.olean /playground diff --git a/lake-manifest.json b/lake-manifest.json index 6014e9c0..c2fc68cd 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -1,60 +1,59 @@ -{"version": 6, - "packagesDir": "lake-packages", +{"version": 7, + "packagesDir": ".lake/packages", "packages": - [{"git": - {"url": "https://github.com/leanprover-community/mathlib4", - "subDir?": null, - "rev": "81dd376a02781030ead59ee35ca5334a7fccc527", - "opts": {}, - "name": "mathlib", - "inputRev?": "81dd376a02781030ead59ee35ca5334a7fccc527", - "inherited": false}}, - {"git": - {"url": "https://github.com/verified-optimization/SciLean", - "subDir?": null, - "rev": "aa0284656e717300b742e81e775d944d185c9f04", - "opts": {}, - "name": "scilean", - "inputRev?": "master", - "inherited": false}}, - {"git": - {"url": "https://github.com/leanprover/std4", - "subDir?": null, - "rev": "869c615eb10130c0637a7bc038e2b80253559913", - "opts": {}, - "name": "std", - "inputRev?": "main", - "inherited": true}}, - {"git": - {"url": "https://github.com/leanprover-community/quote4", - "subDir?": null, - "rev": "511eb96eca98a7474683b8ae55193a7e7c51bc39", - "opts": {}, - "name": "Qq", - "inputRev?": "master", - "inherited": true}}, - {"git": - {"url": "https://github.com/leanprover-community/aesop", - "subDir?": null, - "rev": "cb87803274405db79ec578fc07c4730c093efb90", - "opts": {}, - "name": "aesop", - "inputRev?": "master", - "inherited": true}}, - {"git": - {"url": "https://github.com/leanprover/lean4-cli", - "subDir?": null, - "rev": "a751d21d4b68c999accb6fc5d960538af26ad5ec", - "opts": {}, - "name": "Cli", - "inputRev?": "main", - "inherited": true}}, - {"git": - {"url": "https://github.com/leanprover-community/ProofWidgets4", - "subDir?": null, - "rev": "f1a5c7808b001305ba07d8626f45ee054282f589", - "opts": {}, - "name": "proofwidgets", - "inputRev?": "v0.0.21", - "inherited": true}}], - "name": "CvxLean"} + [{"url": "https://github.com/leanprover/std4", + "type": "git", + "subDir": null, + "rev": "16d8352f7ed0d38cbc58ace03b3429d693cf50c6", + "name": "std", + "manifestFile": "lake-manifest.json", + "inputRev": "main", + "inherited": true, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/leanprover-community/quote4", + "type": "git", + "subDir": null, + "rev": "ccba5d35d07a448fab14c0e391c8105df6e2564c", + "name": "Qq", + "manifestFile": "lake-manifest.json", + "inputRev": "master", + "inherited": true, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/leanprover-community/aesop", + "type": "git", + "subDir": null, + "rev": "3141402ba5a5f0372d2378fd75a481bc79a74ecf", + "name": "aesop", + "manifestFile": "lake-manifest.json", + "inputRev": "master", + "inherited": true, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/leanprover-community/ProofWidgets4", + "type": "git", + "subDir": null, + "rev": "909febc72b4f64628f8d35cd0554f8a90b6e0749", + "name": "proofwidgets", + "manifestFile": "lake-manifest.json", + "inputRev": "v0.0.23", + "inherited": true, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/leanprover/lean4-cli", + "type": "git", + "subDir": null, + "rev": "a751d21d4b68c999accb6fc5d960538af26ad5ec", + "name": "Cli", + "manifestFile": "lake-manifest.json", + "inputRev": "main", + "inherited": true, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/leanprover-community/mathlib4", + "type": "git", + "subDir": null, + "rev": "1250aa83953a2c7d5819cebea08ad7fdef997d49", + "name": "mathlib", + "manifestFile": "lake-manifest.json", + "inputRev": "1250aa83953a2c7d5819cebea08ad7fdef997d49", + "inherited": false, + "configFile": "lakefile.lean"}], + "name": "CvxLean", + "lakeDir": ".lake"} diff --git a/lakefile.lean b/lakefile.lean index 490d476e..75b0c0f4 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -3,8 +3,9 @@ open System Lake DSL require mathlib from git "https://github.com/leanprover-community/mathlib4" @ - "81dd376a02781030ead59ee35ca5334a7fccc527" + "1250aa83953a2c7d5819cebea08ad7fdef997d49" +meta if get_config? env = some "dev" then require scilean from git "https://github.com/verified-optimization/SciLean" @ "master" diff --git a/lean-toolchain b/lean-toolchain index e8560170..91ccf6ac 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.3.0-rc1 +leanprover/lean4:v4.4.0-rc1 From aa205933757f3a58d4a17effe98ea798423e99f8 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 15 Dec 2023 13:31:02 -0500 Subject: [PATCH 186/189] fix: broken proofs from version update --- CvxLean/Lib/Math/CovarianceEstimation.lean | 5 --- CvxLean/Lib/Math/Data/Real.lean | 10 ++--- .../Math/LinearAlgebra/Matrix/ToUpperTri.lean | 2 +- CvxLean/Tactic/Arith/Arith.lean | 42 ------------------- CvxLean/Tactic/Arith/PositivityExt.lean | 4 +- CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean | 7 ++-- CvxLean/Tactic/DCP/AtomLibrary/Fns/Inv.lean | 4 +- CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean | 2 +- .../Tactic/DCP/AtomLibrary/Fns/LogDet.lean | 2 +- CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean | 28 ++++++------- .../DCP/AtomLibrary/Fns/QuadOverLin.lean | 8 ++-- CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean | 18 ++++---- CvxLean/Tactic/DCP/AtomLibrary/Fns/XExp.lean | 6 +-- CvxLean/Tactic/Solver/Float/RealToFloat.lean | 6 +-- 14 files changed, 48 insertions(+), 96 deletions(-) diff --git a/CvxLean/Lib/Math/CovarianceEstimation.lean b/CvxLean/Lib/Math/CovarianceEstimation.lean index 0d5ce259..424d25cb 100644 --- a/CvxLean/Lib/Math/CovarianceEstimation.lean +++ b/CvxLean/Lib/Math/CovarianceEstimation.lean @@ -18,7 +18,6 @@ noncomputable def covarianceMatrix {N n : ℕ} (y : Fin N → Fin n → ℝ) : M lemma gaussianPdf_pos {n : ℕ} (R : Matrix (Fin n) (Fin n) ℝ) (y : Fin n → ℝ) (h : R.PosDef): 0 < gaussianPdf R y := by refine' mul_pos (div_pos zero_lt_one (sqrt_pos.2 (mul_pos _ h.det_pos))) (exp_pos _) - simp [rpow_eq_pow] exact (pow_pos (mul_pos (by positivity) pi_pos) n) lemma prod_gaussianPdf_pos {N n : ℕ} (R : Matrix (Fin n) (Fin n) ℝ) (y : Fin N → Fin n → ℝ) @@ -33,7 +32,6 @@ lemma log_prod_gaussianPdf {N n : ℕ} (y : Fin N → Fin n → ℝ) (R : Matrix i ∈ Finset.univ → gaussianPdf R (y i) ≠ 0 := λ i hi => ne_of_gt (gaussianPdf_pos _ _ hR) have sqrt_2_pi_n_R_det_ne_zero: sqrt ((2 * π) ^ n * R.det) ≠ 0 := by refine' ne_of_gt (sqrt_pos.2 (mul_pos _ hR.det_pos)) - simp [rpow_eq_pow] exact (pow_pos (mul_pos (by positivity) pi_pos) _) rw [log_prod Finset.univ (λ i => gaussianPdf R (y i)) this] unfold gaussianPdf @@ -43,15 +41,12 @@ lemma log_prod_gaussianPdf {N n : ℕ} (y : Fin N → Fin n → ℝ) (R : Matrix simp [rpow_eq_pow] exact ne_of_gt (sqrt_pos.2 (pow_pos (mul_pos (by positivity) pi_pos) _)) exact ne_of_gt (sqrt_pos.2 hR.det_pos) - simp [rpow_eq_pow] exact pow_nonneg (mul_nonneg (by positivity) (le_of_lt pi_pos)) _ norm_num exact sqrt_2_pi_n_R_det_ne_zero exact div_ne_zero (by norm_num) sqrt_2_pi_n_R_det_ne_zero exact exp_ne_zero _ -#check congrArg - lemma sum_quadForm {n : ℕ} (R : Matrix (Fin n) (Fin n) ℝ) {m : ℕ} (y : Fin m → Fin n → ℝ) : (∑ i, R.quadForm (y i)) = m * (covarianceMatrix y * Rᵀ).trace := by diff --git a/CvxLean/Lib/Math/Data/Real.lean b/CvxLean/Lib/Math/Data/Real.lean index 683d6d29..b4106383 100644 --- a/CvxLean/Lib/Math/Data/Real.lean +++ b/CvxLean/Lib/Math/Data/Real.lean @@ -56,23 +56,23 @@ lemma div_pow_eq_mul_pow_neg {a b c : ℝ} (hb : 0 ≤ b) : a / (b ^ c) = a * b ^ (-c) := by rw [div_eq_mul_inv, ←rpow_neg hb] -lemma one_div_eq_pow_neg_one {a : ℝ} (ha : 0 < a) : 1 / a = a ^ (-1) := by +lemma one_div_eq_pow_neg_one {a : ℝ} (ha : 0 < a) : 1 / a = a ^ (-1 : ℝ) := by rw [rpow_neg (le_of_lt ha), rpow_one, div_eq_mul_inv, one_mul] -lemma inv_eq_pow_neg_one {a : ℝ} (ha : 0 < a) : a⁻¹ = a ^ (-1) := by +lemma inv_eq_pow_neg_one {a : ℝ} (ha : 0 < a) : a⁻¹ = a ^ (-1 : ℝ) := by rw [inv_eq_one_div, one_div_eq_pow_neg_one ha] -lemma pow_half_two {x : ℝ} (hx : 0 ≤ x) : (x ^ (1 / 2)) ^ 2 = x := by +lemma pow_half_two {x : ℝ} (hx : 0 ≤ x) : (x ^ (1 / 2 : ℝ)) ^ (2 : ℝ) = x := by show rpow (rpow _ _) _ = _ rw [rpow_eq_pow, rpow_eq_pow, ← rpow_mul hx] norm_num lemma pow_two_le_pow_two {x y : ℝ} (hx : 0 ≤ x) (hy : 0 ≤ y) : - x ^ 2 ≤ y ^ 2 ↔ x ≤ y := by + x ^ (2 : ℝ) ≤ y ^ (2 : ℝ) ↔ x ≤ y := by rw [rpow_two, rpow_two, sq_le_sq, abs_of_nonneg hx, abs_of_nonneg hy] lemma binomial_two (x y : ℝ) : - (x + y) ^ 2 = x ^ 2 + (2 * (x * y) + y ^ 2) := by + (x + y) ^ (2 : ℝ) = x ^ (2 : ℝ) + (2 * (x * y) + y ^ (2 : ℝ)) := by simp only [rpow_two]; ring lemma rpow_eq_mul_rpow_pred {x y : ℝ} (hx : x ≠ 0) : diff --git a/CvxLean/Lib/Math/LinearAlgebra/Matrix/ToUpperTri.lean b/CvxLean/Lib/Math/LinearAlgebra/Matrix/ToUpperTri.lean index af4780fd..80f0f200 100644 --- a/CvxLean/Lib/Math/LinearAlgebra/Matrix/ToUpperTri.lean +++ b/CvxLean/Lib/Math/LinearAlgebra/Matrix/ToUpperTri.lean @@ -36,7 +36,7 @@ lemma upperTriangular_toUpperTri {m : Type _} [LinearOrder m] (A : Matrix m m lemma upperTriangular.toUpperTri_eq {A : Matrix n n ℝ} (hA : upperTriangular A) : A.toUpperTri = A := by ext i j - by_cases i ≤ j + by_cases h : i ≤ j simp [toUpperTri, h] simp [toUpperTri, h, hA (lt_of_not_ge h)] diff --git a/CvxLean/Tactic/Arith/Arith.lean b/CvxLean/Tactic/Arith/Arith.lean index 39e5ab0d..69dc5f96 100644 --- a/CvxLean/Tactic/Arith/Arith.lean +++ b/CvxLean/Tactic/Arith/Arith.lean @@ -102,45 +102,3 @@ syntax "arith" : tactic macro_rules | `(tactic| arith) => `(tactic| (first | linarith | positivity! | norm_num_simp_pow)) - -open Real - -lemma xxx : ∀ (x : ℝ × ℝ × ℝ × ℝ), - log (OfScientific.ofScientific 10 true 0) ≤ - x.2.2.1 - x.1 - - OfScientific.ofScientific 5 true 1 * - log (rexp (OfScientific.ofScientific 2 true 0 * x.1) + rexp (OfScientific.ofScientific 2 true 0 * x.2.1)) → - x.2.1 ≤ - x.2.2.1 - - (OfScientific.ofScientific 5 true 1 * - log - (rexp (OfScientific.ofScientific 2 true 0 * x.1) + - rexp (OfScientific.ofScientific 2 true 0 * x.2.1)) + - log (OfScientific.ofScientific 20 true 0)) → - OfScientific.ofScientific 0 true 0 ≤ x.1 → - x.1 ≤ log (OfScientific.ofScientific 100 true 0) → - OfScientific.ofScientific 0 true 0 ≤ x.2.1 → - x.2.1 ≤ log (OfScientific.ofScientific 100 true 0) → - sqrt (rexp x.2.2.1 / (314159 / 50000) + rexp x.2.2.2 ^ 2) ≤ 10 → - (log - (rexp x.2.2.2 ^ OfScientific.ofScientific 2 true 0 * - ((OfScientific.ofScientific 11 true 0 / OfScientific.ofScientific 10 true 0) ^ - OfScientific.ofScientific 2 true 0 - - OfScientific.ofScientific 1 true 0)) ≤ - log - (rexp x.2.2.1 / - (OfScientific.ofScientific 314159 true 0 / OfScientific.ofScientific 50000 true 0)) ↔ - log (rexp x.2.2.2 ^ OfScientific.ofScientific 2 true 0) + - log - ((OfScientific.ofScientific 11 true 0 / OfScientific.ofScientific 10 true 0) ^ - OfScientific.ofScientific 2 true 0 - - OfScientific.ofScientific 1 true 0) ≤ - log - (rexp x.2.2.1 / - (OfScientific.ofScientific 314159 true 0 / OfScientific.ofScientific 50000 true 0))) := by { - norm_num_clean_up - intros - rw [log_mul] - { positivity! } - { norm_cast; positivity! } - } diff --git a/CvxLean/Tactic/Arith/PositivityExt.lean b/CvxLean/Tactic/Arith/PositivityExt.lean index d63b3c8a..32bea574 100644 --- a/CvxLean/Tactic/Arith/PositivityExt.lean +++ b/CvxLean/Tactic/Arith/PositivityExt.lean @@ -108,12 +108,12 @@ def evalOneSubDivExp : PositivityExt where eval {_ _α} zα pα e := do pure .none lemma Real.scaled_sq_diff_pos_of_pos {a x : ℝ} : - 0 < a - 1 → 0 < x → 0 < (a * x) ^ 2 - x ^ 2 := + 0 < a - 1 → 0 < x → 0 < (a * x) ^ (2 : ℝ) - x ^ (2 : ℝ) := fun h1 h2 => by have ha : 0 ≤ a := by linarith have ha1 : 1 < a := by linarith have hx : 0 ≤ x := by linarith - have hx2 : 0 < x ^ 2 := by positivity + have hx2 : 0 < x ^ (2 : ℝ) := by positivity rw [sub_pos, Real.mul_rpow ha hx, ←div_lt_iff hx2, div_self (ne_of_gt hx2)] simp [ha1, abs_eq_self.mpr ha] diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean index 91ddf5b9..a8eead7d 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Huber.lean @@ -17,7 +17,7 @@ open Real declare_atom huber [convex] (x : ℝ)? : huber x := vconditions implementationVars (v : ℝ) (w : ℝ) -implementationObjective (2 * v + w ^ 2) +implementationObjective (2 * v + w ^ (2 : ℝ)) implementationConstraints (c1 : |x| ≤ v + w) (c2 : 0 ≤ w) @@ -42,7 +42,8 @@ optimality by simp [huber] split_ifs with h { by_cases hwx : w ≤ |x| - { have hsq : |x| ^ 2 - w ^ 2 = (|x| + w) * (|x| - w) := by ring_nf; simp + { have hsq : |x| ^ (2 : ℝ) - w ^ (2 : ℝ) = (|x| + w) * (|x| - w) := by + ring_nf; simp rw [←sub_le_iff_le_add, ←sq_abs, ←rpow_two, ←rpow_two, hsq] apply mul_le_mul <;> linarith } { replace hwx := not_le.mp hwx @@ -67,7 +68,7 @@ vconditionElimination declare_atom Vec.huber [convex] (m : Nat)& (x : Fin m → ℝ)? : Vec.huber x := vconditions implementationVars (v : Fin m → ℝ) (w : Fin m → ℝ) -implementationObjective (2 • v + w ^ 2) +implementationObjective (2 • v + w ^ (2 : ℝ)) implementationConstraints (c1 : |x| ≤ v + w) (c2 : 0 ≤ w) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Inv.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Inv.lean index 1c14340f..bcfae04c 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Inv.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Inv.lean @@ -33,7 +33,7 @@ feasibility exact powNegOne.feasibility0 x hx) (c2 : by rw [Real.inv_eq_pow_neg_one hx] - have hinv : 0 < x ^ (-1) := by rwa [rpow_neg_one, inv_pos] + have hinv : 0 < x ^ (-1 : ℝ) := by rwa [rpow_neg_one, inv_pos] exact le_of_lt hinv) (c3 : le_of_lt hx) optimality by @@ -74,7 +74,7 @@ feasibility exact powNegOne.feasibility0 x hx) (c2 : by rw [Real.one_div_eq_pow_neg_one hx] - have hinv : 0 < x ^ (-1) := by rwa [rpow_neg_one, inv_pos] + have hinv : 0 < x ^ (-1 : ℝ) := by rwa [rpow_neg_one, inv_pos] exact le_of_lt hinv) (c3 : le_of_lt hx) optimality by diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean index fb84f360..c877534d 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/KLDiv.lean @@ -116,7 +116,7 @@ implementationConstraints (c2 : Vec.const m (1 / 100000) ≤ y) solution (t := fun i => -((x i) * log ((x i) / (y i)))) solutionEqualsAtom by - simp [Vec.klDiv, klDiv]; ext i; simp; ring + unfold Vec.klDiv klDiv; ext i; simp; ring feasibility (c1 : by simp [Vec.klDiv, klDiv] diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogDet.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogDet.lean index ed6a46f9..08697222 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogDet.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/LogDet.lean @@ -37,7 +37,7 @@ solution if h : A.PosDef then (LDL.diag h) * (LDL.lower h)ᵀ else 0) solutionEqualsAtom by simp only [dif_pos hA, Vec.sum, Vec.log] - rw [Matrix.LogDetAtom.solution_eq_atom hA] + erw [Matrix.LogDetAtom.solution_eq_atom hA] congr feasibility (c_exp : by diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean index c967b899..9fe3b499 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Power.lean @@ -20,7 +20,7 @@ optimality by intros _ h simp [h] -declare_atom powNegOne [convex] (x : ℝ)- : x ^ (-1) := +declare_atom powNegOne [convex] (x : ℝ)- : x ^ (-1 : ℝ) := vconditions (hx : 0 < x) implementationVars (t : ℝ) @@ -30,17 +30,17 @@ implementationConstraints (c2 : 0 ≤ t) (c3 : 0 ≤ x) solution - (t := x ^ (-1)) + (t := x ^ (-1 : ℝ)) solutionEqualsAtom rfl feasibility (c1 : by have hxnn := le_of_lt hx - have hinv : 0 < x ^ (-1) := by rwa [rpow_neg_one, inv_pos] + have hinv : 0 < x ^ (-1 : ℝ) := by rwa [rpow_neg_one, inv_pos] have hinvnn := le_of_lt hinv rw [soCone_add_sub_two_of_nonneg hinvnn hxnn, rpow_neg_one] field_simp) (c2 : by - have hinv : 0 < x ^ (-1) := by rwa [rpow_neg_one, inv_pos] + have hinv : 0 < x ^ (-1 : ℝ) := by rwa [rpow_neg_one, inv_pos] exact le_of_lt hinv) (c3 : le_of_lt hx) optimality by @@ -65,7 +65,7 @@ vconditionElimination -- NOTE(RFM): It is not straightforward to express it in terms of x ^ (-1) and -- x ^ 2 because of vconditions. -declare_atom powNegTwo [convex] (x : ℝ)- : x ^ (-2) := +declare_atom powNegTwo [convex] (x : ℝ)- : x ^ (-2 : ℝ) := vconditions (hx : 0 < x) implementationVars (t₀ : ℝ) (t₁ : ℝ) @@ -77,21 +77,21 @@ implementationConstraints (c4 : 0 ≤ t₁) (c5 : 0 ≤ x) solution - (t₀ := x ^ (-2)) (t₁ := x ^ (-1)) + (t₀ := x ^ (-2 : ℝ)) (t₁ := x ^ (-1 : ℝ)) solutionEqualsAtom rfl feasibility (c1 : by dsimp have hxnn := le_of_lt hx - have hinv : 0 < x ^ (-1) := by rwa [rpow_neg_one, inv_pos] + have hinv : 0 < x ^ (-1 : ℝ) := by rwa [rpow_neg_one, inv_pos] have hinvnn := le_of_lt hinv rw [soCone_add_sub_two_of_nonneg hxnn hinvnn, rpow_neg_one] field_simp) (c2 : by dsimp have hxnn := le_of_lt hx - have hinv : 0 < x ^ (-1) := by rwa [rpow_neg_one, inv_pos] - have hxneg2 : 0 < x ^ (-2) := by + have hinv : 0 < x ^ (-1 : ℝ) := by rwa [rpow_neg_one, inv_pos] + have hxneg2 : 0 < x ^ (-2 : ℝ) := by rw [(by norm_num : (-2 : ℝ) = (-1) + (-1)), rpow_add hx] apply mul_pos <;> exact hinv have hxneg2nn := le_of_lt hxneg2 @@ -130,24 +130,24 @@ optimality by have ht₁2pos : 0 < t₁ ^ (2 : ℝ) := by rw [rpow_two] exact pow_two_pos_of_ne_zero t₁ (ne_of_gt ht₁pos) - have ht₁inv : 0 < t₁ ^ (-1) := by rwa [rpow_neg_one, inv_pos] + have ht₁inv : 0 < t₁ ^ (-1 : ℝ) := by rwa [rpow_neg_one, inv_pos] have ht₀pos : 0 < t₀ := by cases (lt_or_eq_of_le c3) with | inl h => exact h | inr h => rw [←h] at c2; linarith -- Combine c1 and c2 appropriately to get t₀ ^ (-1) ≤ y ^ 2. - have ht₁invx : t₁ ^ (-1) ≤ x := by + have ht₁invx : t₁ ^ (-1 : ℝ) ≤ x := by rw [rpow_neg c4, rpow_one] rwa [←div_le_iff ht₁pos, ←inv_eq_one_div] at c1 - have ht₁invy : t₁ ^ (-1) ≤ y := le_trans ht₁invx hy + have ht₁invy : t₁ ^ (-1 : ℝ) ≤ y := le_trans ht₁invx hy have ht₁invnn := le_of_lt ht₁inv - have ht₁neg2y2 : t₁ ^ (-2) ≤ y ^ (2 : ℝ) := by + have ht₁neg2y2 : t₁ ^ (-2 : ℝ) ≤ y ^ (2 : ℝ) := by have h := mul_le_mul ht₁invy ht₁invy ht₁invnn hynn rw [←rpow_add ht₁pos] at h norm_num at h rw [←pow_two, ←rpow_two] at h exact h - have ht₀invt₁neg2 : t₀ ^ (-1) ≤ t₁ ^ (-2) := by + have ht₀invt₁neg2 : t₀ ^ (-1 : ℝ) ≤ t₁ ^ (-2 : ℝ) := by rw [mul_one] at c2 have ht₀1 : 0 < t₀ ^ (1 : ℝ) := by rwa [rpow_one] rw [rpow_neg c3, rpow_neg c4, inv_le_inv ht₀1 ht₁2pos, rpow_one] diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean index ff3276eb..33a1077e 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/QuadOverLin.lean @@ -15,22 +15,22 @@ open Real -- TODO(RFM): generalize to x : Fin n → ℝ -- TODO(RFM): distinguish between nonincreasing and nondecreasing. -- TODO(RFM): how to do strict positivity and make vcondelimination work? -declare_atom quadOverLin [convex] (x : ℝ)? (y : ℝ)- : x ^ 2 / y := +declare_atom quadOverLin [convex] (x : ℝ)? (y : ℝ)- : x ^ (2 : ℝ) / y := vconditions (hy : 1 / 100000 ≤ y) -implementationVars (t : ℝ) (y' : ℝ) +implementationVars (t : ℝ) implementationObjective (t) implementationConstraints (c1 : soCone (y + t) ![y - t, 2 * x]) (c2 : 0 ≤ t) (c3 : 1 / 100000 ≤ y) -solution (t := x ^ 2 / y) (y' := log y) +solution (t := x ^ (2 : ℝ) / y) solutionEqualsAtom by rfl feasibility (c1 : by have hypos : 0 < y := by positivity have hynn : 0 ≤ y := by linarith - have hx2ynn : 0 ≤ x ^ 2 / y := by + have hx2ynn : 0 ≤ x ^ (2 : ℝ) / y := by rw [rpow_two]; exact div_nonneg (pow_two_nonneg x) hynn rw [soCone_add_sub_two_mul_of_nonneg _ hynn hx2ynn] rw [mul_div, mul_comm, ←mul_div, div_self (ne_of_gt hypos), mul_one]) diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean index 09cf5094..35e53edf 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/Sq.lean @@ -7,14 +7,14 @@ namespace CvxLean open Real -declare_atom sq [convex] (x : ℝ)? : x ^ 2 := +declare_atom sq [convex] (x : ℝ)? : x ^ (2 : ℝ) := vconditions implementationVars (t : ℝ) implementationObjective (t) implementationConstraints - (c1 : rotatedSoCone t (1/2) ![x]) + (c1 : rotatedSoCone t (1 / 2) ![x]) solution - (t := x ^ 2) + (t := x ^ (2 : ℝ)) solutionEqualsAtom rfl feasibility (c1 : by @@ -28,24 +28,24 @@ optimality by exact h vconditionElimination -declare_atom Vec.sq [convex] (n : ℕ)& (x : Fin n → ℝ)? : x ^ 2 := +declare_atom Vec.sq [convex] (n : ℕ)& (x : Fin n → ℝ)? : x ^ (2 : ℝ) := vconditions implementationVars (t : Fin n → ℝ) implementationObjective (t) implementationConstraints - (c1 : Vec.rotatedSoCone t (fun _ => 1/2) (Vec.toMatrix x)) + (c1 : Vec.rotatedSoCone t (fun _ => 1 / 2) (Vec.toMatrix x)) solution - (t := x ^ 2) + (t := x ^ (2 : ℝ)) solutionEqualsAtom rfl feasibility (c1 : by unfold Vec.rotatedSoCone - intros t i - convert sq.feasibility0 (x i); simp) + intros _ i + exact sq.feasibility0 (x i)) optimality by intros i unfold Vec.rotatedSoCone at c1 - convert sq.optimality (x i) (t i) (c1 i); simp + exact sq.optimality (x i) (t i) (c1 i) vconditionElimination end CvxLean diff --git a/CvxLean/Tactic/DCP/AtomLibrary/Fns/XExp.lean b/CvxLean/Tactic/DCP/AtomLibrary/Fns/XExp.lean index 2a6fda06..f7a317b7 100644 --- a/CvxLean/Tactic/DCP/AtomLibrary/Fns/XExp.lean +++ b/CvxLean/Tactic/DCP/AtomLibrary/Fns/XExp.lean @@ -12,14 +12,14 @@ implementationVars (t₀ : ℝ) (t₁ : ℝ) implementationObjective t₀ implementationConstraints (c1 : expCone t₁ x t₀) - (c2 : x ^ 2 ≤ t₁) + (c2 : x ^ (2 : ℝ) ≤ t₁) (c3 : 0 ≤ x) -solution (t₀ := x * exp x) (t₁ := x ^ 2) +solution (t₀ := x * exp x) (t₁ := x ^ (2 : ℝ)) solutionEqualsAtom rfl feasibility (c1 : by unfold expCone - by_cases (0 < x) + by_cases h : 0 < x . left refine ⟨h, ?_⟩ rw [rpow_two, pow_two, ←mul_div, div_self (ne_of_lt h).symm, mul_one] diff --git a/CvxLean/Tactic/Solver/Float/RealToFloat.lean b/CvxLean/Tactic/Solver/Float/RealToFloat.lean index bbb4314f..abdb9fb9 100644 --- a/CvxLean/Tactic/Solver/Float/RealToFloat.lean +++ b/CvxLean/Tactic/Solver/Float/RealToFloat.lean @@ -130,7 +130,7 @@ addRealToFloat (x : ℕ) (i) : @instOfNat Real x Real.natCast i := @instOfNatFloat x addRealToFloat (k : Nat) : - @SMul.smul ℕ ℝ AddMonoid.SMul k := + @SMul.smul ℕ ℝ AddMonoid.toNatSMul k := (fun (x : Float) => (OfNat.ofNat k) * x) addRealToFloat (i) : @Ring.toNeg Real i := @@ -208,7 +208,7 @@ addRealToFloat (n m k) : (fun (M : Matrix (Fin n) (Fin m) Float) i j => (OfNat.ofNat k) * (M i j)) addRealToFloat (n m k : Nat) : - @SMul.smul ℕ (Matrix (Fin n) (Fin m) ℝ) AddMonoid.SMul k := + @SMul.smul ℕ (Matrix (Fin n) (Fin m) ℝ) AddMonoid.toNatSMul k := (fun (M : Matrix (Fin n) (Fin m) Float) i j => (OfNat.ofNat k) * (M i j)) addRealToFloat : @Matrix.vecEmpty Real := @@ -248,8 +248,6 @@ addRealToFloat (n : Nat) (i) : @Matrix.trace (Fin n) ℝ (Fin.fintype n) i := @Matrix.Computable.tr Float (Zero.mk (0 : Float)) n instAddFloat -#check HMul.hMul - addRealToFloat (n : Nat) (i) : @HMul.hMul (Matrix (Fin n) (Fin n) ℝ) From 2507d04cbd5e7a67a68226c00ef762b9cf615f8f Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 15 Dec 2023 15:12:14 -0500 Subject: [PATCH 187/189] fix: tests passing with new Lean version --- CvxLean.lean | 1 + CvxLean/Tactic/Arith/PositivityExt.lean | 4 +- CvxLean/Tactic/ChangeOfVariables/Basic.lean | 2 +- CvxLean/Tactic/Convexify/MapObjFun.lean | 4 +- CvxLean/Test/Convexify/AlmostDGP.lean | 4 +- CvxLean/Test/Convexify/DGP.lean | 20 +++++----- CvxLean/Test/Convexify/DQCP.lean | 4 +- CvxLean/Test/Convexify/Quiz.lean | 10 ++--- CvxLean/Test/Convexify/Unit.lean | 42 ++++++++++----------- CvxLean/Test/DCP/Dcp.lean | 2 +- 10 files changed, 47 insertions(+), 46 deletions(-) diff --git a/CvxLean.lean b/CvxLean.lean index 15948caf..51203208 100644 --- a/CvxLean.lean +++ b/CvxLean.lean @@ -1 +1,2 @@ import CvxLean.Command.Solve +import CvxLean.Tactic.Convexify.Convexify diff --git a/CvxLean/Tactic/Arith/PositivityExt.lean b/CvxLean/Tactic/Arith/PositivityExt.lean index 32bea574..a0ad9820 100644 --- a/CvxLean/Tactic/Arith/PositivityExt.lean +++ b/CvxLean/Tactic/Arith/PositivityExt.lean @@ -9,11 +9,11 @@ namespace Mathlib.Meta.Positivity open Lean.Meta Qq lemma Real.one_sub_one_div_sq_nonneg_of_le_one {x : ℝ} : - 0 < x → 0 ≤ 1 - x → 0 ≤ (1 / x ^ 2) - 1 := + 0 < x → 0 ≤ 1 - x → 0 ≤ (1 / x ^ (2 : ℝ)) - 1 := fun h1 h2 => by have hx1 : x ≤ 1 := by linarith have h0x : 0 ≤ x := by positivity - have h0x2 : 0 < x ^ 2 := by positivity + have h0x2 : 0 < x ^ (2 : ℝ) := by positivity rw [le_sub_iff_add_le, zero_add, le_div_iff h0x2, one_mul] simpa [abs_eq_self.mpr h0x] diff --git a/CvxLean/Tactic/ChangeOfVariables/Basic.lean b/CvxLean/Tactic/ChangeOfVariables/Basic.lean index 02a9c57a..711510e5 100644 --- a/CvxLean/Tactic/ChangeOfVariables/Basic.lean +++ b/CvxLean/Tactic/ChangeOfVariables/Basic.lean @@ -125,7 +125,7 @@ macro "map_exp" : tactic => (g := fun x => ExpMap.exp x) (f := fun x => LogMap.log x) (hfg := by prove_exp_log) <;> - dsimp only [Function.comp, ExpMap.exp, LogMap.log] <;> + simp only [Function.comp, ExpMap.exp, LogMap.log] <;> remove_positive_constraints) /-- Same as `map_exp` but at a particular position in the domain product. -/ diff --git a/CvxLean/Tactic/Convexify/MapObjFun.lean b/CvxLean/Tactic/Convexify/MapObjFun.lean index 137a87f1..62563503 100644 --- a/CvxLean/Tactic/Convexify/MapObjFun.lean +++ b/CvxLean/Tactic/Convexify/MapObjFun.lean @@ -29,7 +29,7 @@ elab (name := prove_log_le_log) "prove_log_le_log" : tactic => do macro "map_objFun_log" : tactic => `(tactic| apply map_objective (g := Real.log) (hg := by prove_log_le_log) <;> - dsimp only [Function.comp]) + simp only [Function.comp]) /-- Tactic `map_objFun_sq` used to square the objective function attempting to prove all the side conditions with simple tactics. -/ @@ -47,7 +47,7 @@ elab (name := prove_pow_two_le_pow_two) "prove_pow_two_le_pow_two" : tactic => d macro "map_objFun_sq" : tactic => `(tactic| apply map_objective (g := fun x => x ^ (2 : ℝ)) (hg := by prove_pow_two_le_pow_two) <;> - dsimp only [Function.comp]) + simp only [Function.comp]) end Tactic diff --git a/CvxLean/Test/Convexify/AlmostDGP.lean b/CvxLean/Test/Convexify/AlmostDGP.lean index 4c49d69d..1e24b93b 100644 --- a/CvxLean/Test/Convexify/AlmostDGP.lean +++ b/CvxLean/Test/Convexify/AlmostDGP.lean @@ -17,7 +17,7 @@ def agp1 := minimize (x) subject to h1 : 0 < x - h2 : x ^ 2 - 10.123 ≤ 0 + h2 : x ^ (2 : ℝ) - 10.123 ≤ 0 time_cmd reduction reda1/dcpa1 : agp1 := by map_exp @@ -72,7 +72,7 @@ def agp3 := h3 : 0 < z h4 : 2 ≤ x h5 : x ≤ 3 - h6 : sqrt x ≤ x ^ 2 - 6 * y / z + h6 : sqrt x ≤ x ^ (2 : ℝ) - 6 * y / z h7 : x * y = z time_cmd reduction reda3/dcpa3 : agp3 := by diff --git a/CvxLean/Test/Convexify/DGP.lean b/CvxLean/Test/Convexify/DGP.lean index 9294f454..08551bd2 100644 --- a/CvxLean/Test/Convexify/DGP.lean +++ b/CvxLean/Test/Convexify/DGP.lean @@ -16,7 +16,7 @@ def gp1 := minimize (x) subject to h1 : 0 < x - h2 : x ^ 2 ≤ 10.123 + h2 : x ^ (2 : ℝ) ≤ 10.123 time_cmd reduction red1/dcp1 : gp1 := by map_exp @@ -94,7 +94,7 @@ def gp4 := h3 : 0 < z h4 : 2 <= x h5 : x <= 3 - h6 : x ^ 2 + 6 * y / z <= sqrt x + h6 : x ^ (2 : ℝ) + 6 * y / z <= sqrt x h7 : x * y = z time_cmd reduction red4/dcp4 : gp4 := by @@ -128,8 +128,8 @@ def gp5 := h3 : 0 < z h4 : 2 ≤ x h5 : x ≤ 3 - h6 : x^2 + 3 * y / z ≤ sqrt x - h7 : x / y = z ^ 2 + h6 : x ^ (2 : ℝ) + 3 * y / z ≤ sqrt x + h7 : x / y = z ^ (2 : ℝ) time_cmd reduction red5/dcp5 : gp5 := by map_exp @@ -202,7 +202,7 @@ def gp7 := h1 : 0 < x h2 : 0 < y h3 : 0 < z - h4 : (1 / 3) * (1 / (x ^ 2)) * (1 / (y ^ 2)) + (4 / 3) * (sqrt y) * (1 / z) ≤ 1 + h4 : (1 / 3) * (1 / (x ^ (2 : ℝ))) * (1 / (y ^ (2 : ℝ))) + (4 / 3) * (sqrt y) * (1 / z) ≤ 1 h5 : x + 2 * y + 3 * z ≤ 1 h6 : (1 / 2) * x * y = 1 @@ -230,20 +230,20 @@ section GP8 -- hmin = wmin = 1, hmax = wmax = 100, Rmax = 10, σ = 0.5, π ≈ 3.14159. def gp8 := optimization (h w A r : ℝ) - minimize 2 * A * sqrt (w ^ 2 + h ^ 2) + minimize 2 * A * sqrt (w ^ (2 : ℝ) + h ^ (2 : ℝ)) subject to h1 : 0 < h h2 : 0 < w h3 : 0 < A h4 : 0 < r - h5 : 10 * sqrt (w ^ 2 + h ^ 2) / 2 * h ≤ 0.5 * A - h6 : 20 * sqrt (w ^ 2 + h ^ 2) / 2 * w ≤ 0.5 * A + h5 : 10 * sqrt (w ^ (2 : ℝ) + h ^ (2 : ℝ)) / 2 * h ≤ 0.5 * A + h6 : 20 * sqrt (w ^ (2 : ℝ) + h ^ (2 : ℝ)) / 2 * w ≤ 0.5 * A h7 : 1 ≤ h h8 : h ≤ 100 h9 : 1 ≤ w h10 : w ≤ 100 - h11 : 1.1 * r ≤ sqrt (A / (2 * 3.14159) + r ^ 2) - h12 : sqrt (A / (2 * 3.14159) + r ^ 2) ≤ 10 + h11 : 1.1 * r ≤ sqrt (A / (2 * 3.14159) + r ^ (2 : ℝ)) + h12 : sqrt (A / (2 * 3.14159) + r ^ (2 : ℝ)) ≤ 10 set_option maxHeartbeats 1000000 in time_cmd reduction red8/dcp8 : gp8 := by diff --git a/CvxLean/Test/Convexify/DQCP.lean b/CvxLean/Test/Convexify/DQCP.lean index 63e97c87..050e53b3 100644 --- a/CvxLean/Test/Convexify/DQCP.lean +++ b/CvxLean/Test/Convexify/DQCP.lean @@ -45,11 +45,11 @@ section QCP2 def hypersonicShapeDesign (a b : ℝ) := optimization (Δx : ℝ) - minimize sqrt (1 / (Δx ^ 2) - 1) + minimize sqrt (1 / (Δx ^ (2 : ℝ)) - 1) subject to h1 : 10e-6 ≤ Δx h2 : Δx ≤ 1 - h3 : a * (1 / Δx) - (1 - b) * sqrt (1 - Δx ^ 2) ≤ 0 + h3 : a * (1 / Δx) - (1 - b) * sqrt (1 - Δx ^ (2 : ℝ)) ≤ 0 time_cmd equivalence redqcp2/dqcp2 : hypersonicShapeDesign 0.05 0.65 := by unfold hypersonicShapeDesign diff --git a/CvxLean/Test/Convexify/Quiz.lean b/CvxLean/Test/Convexify/Quiz.lean index a112669f..dfd0bbe8 100644 --- a/CvxLean/Test/Convexify/Quiz.lean +++ b/CvxLean/Test/Convexify/Quiz.lean @@ -30,7 +30,7 @@ time_cmd equivalence eq2/dcpquiz2 : quiz2 := by def quiz3 := optimization (x : ℝ) - minimize ((sqrt x) ^ 2) + minimize ((sqrt x) ^ (2 : ℝ)) subject to h : 0 ≤ x @@ -57,7 +57,7 @@ time_cmd equivalence eq5/dcpquiz5 : quiz5 := by def quiz6 := optimization (x : ℝ) - minimize (-(log ((364 * x) ^ 2))) + minimize (-(log ((364 * x) ^ (2 : ℝ)))) subject to h : 0 ≤ x @@ -66,7 +66,7 @@ time_cmd equivalence eq6/dcpquiz6 : quiz6 := by def quiz7 := optimization (x : ℝ) - minimize ((sqrt ((x + 2) * (1 / x))) ^ 2) + minimize ((sqrt ((x + 2) * (1 / x))) ^ (2 : ℝ)) subject to h : 0 < x @@ -84,7 +84,7 @@ time_cmd equivalence eq8/dcpquiz8 : quiz8 := by def quiz9 := optimization (x y : ℝ) - minimize (1 / (((x⁻¹) ^ 2) / (y⁻¹))) + minimize (1 / (((x⁻¹) ^ (2 : ℝ)) / (y⁻¹))) subject to h₁ : 0 < x h₂ : 0 < y @@ -94,7 +94,7 @@ time_cmd equivalence eq9/dcpquiz9 : quiz9 := by def quiz10 := optimization (x : ℝ) - minimize ((log (exp x)) ^ 2) + minimize ((log (exp x)) ^ (2 : ℝ)) subject to h : 0 ≤ x diff --git a/CvxLean/Test/Convexify/Unit.lean b/CvxLean/Test/Convexify/Unit.lean index 043d152e..29988d52 100644 --- a/CvxLean/Test/Convexify/Unit.lean +++ b/CvxLean/Test/Convexify/Unit.lean @@ -164,7 +164,7 @@ def powTwoLePowTwoConstr := subject to hx : 0 ≤ x hy : 0 ≤ y - h : x ^ 2 ≤ y ^ 2 + h : x ^ (2 : ℝ) ≤ y ^ (2 : ℝ) time_cmd equivalence powTwoLePowTwoConstrRed/powTwoLePowTwoConstrAuto : powTwoLePowTwoConstr := by convexify @@ -842,7 +842,7 @@ time_cmd equivalence powAddRevConstrRed/powAddRevConstrAuto : powAddRevConstr := -- mul_pow (obj) def mulPowObj := optimization (x : ℝ) - minimize ((x * (sqrt x)) ^ 2 / x : ℝ) + minimize ((x * (sqrt x)) ^ (2 : ℝ) / x : ℝ) subject to hx : 0 < x @@ -857,7 +857,7 @@ def mulPowConstr := minimize (0 : ℝ) subject to hx : 0 < x - h : ((x * (sqrt x)) ^ 2 / x) ≤ 1 + h : ((x * (sqrt x)) ^ (2 : ℝ) / x) ≤ 1 time_cmd equivalence mulPowConstrRed/mulPowConstrAuto : mulPowConstr := by convexify @@ -867,7 +867,7 @@ time_cmd equivalence mulPowConstrRed/mulPowConstrAuto : mulPowConstr := by -- mul_pow-rev (obj) def mulPowRevObj := optimization (x : ℝ) - minimize (((sqrt x) ^ 2) * ((sqrt (x + 1)) ^ 2) : ℝ) + minimize (((sqrt x) ^ (2 : ℝ)) * ((sqrt (x + 1)) ^ (2 : ℝ)) : ℝ) subject to hx : 0 < x @@ -882,7 +882,7 @@ def mulPowRevConstr := minimize (0 : ℝ) subject to hx : 0 < x - h : (((sqrt x) ^ 2) * ((sqrt (x + 1)) ^ 2)) ≤ 1 + h : (((sqrt x) ^ (2 : ℝ)) * ((sqrt (x + 1)) ^ (2 : ℝ))) ≤ 1 time_cmd equivalence mulPowRevConstrRed/mulPowRevConstrAuto : mulPowRevConstr := by convexify @@ -892,7 +892,7 @@ time_cmd equivalence mulPowRevConstrRed/mulPowRevConstrAuto : mulPowRevConstr := -- pow_mul-rev (obj) def powMulRevObj := optimization (x : ℝ) - minimize ((x ^ 1) ^ 2 : ℝ) + minimize ((x ^ (1 : ℝ)) ^ (2 : ℝ) : ℝ) subject to hx : 0 < x @@ -907,7 +907,7 @@ def powMulRevConstr := minimize (0 : ℝ) subject to hx : 0 < x - h : ((x ^ 2) ^ 2) ≤ 1 + h : ((x ^ (2 : ℝ)) ^ (2 : ℝ)) ≤ 1 time_cmd equivalence powMulRevConstrRed/powMulRevConstrAuto : powMulRevConstr := by convexify @@ -917,7 +917,7 @@ time_cmd equivalence powMulRevConstrRed/powMulRevConstrAuto : powMulRevConstr := -- div_pow (obj) def divPowObj := optimization (x : ℝ) - minimize ((x ^ 2) * (1 / x) ^ 2 : ℝ) + minimize ((x ^ (2 : ℝ)) * (1 / x) ^ (2 : ℝ) : ℝ) subject to hx : 0 < x @@ -932,7 +932,7 @@ def divPowConstr := minimize (0 : ℝ) subject to hx : 0 < x - h : (1 / x) ^ 2 ≤ 1 + h : (1 / x) ^ (2 : ℝ) ≤ 1 time_cmd equivalence divPowConstrRed/divPowConstrAuto : divPowConstr := by convexify @@ -942,7 +942,7 @@ time_cmd equivalence divPowConstrRed/divPowConstrAuto : divPowConstr := by -- div_pow-rev (obj) def divPowRevObj := optimization (x : ℝ) - minimize ((x + x) ^ 2 / x ^ 2 : ℝ) + minimize ((x + x) ^ (2 : ℝ) / x ^ (2 : ℝ) : ℝ) subject to hx : 0 < x @@ -957,7 +957,7 @@ def divPowRevConstr := minimize (0 : ℝ) subject to hx : 0 < x - h : ((x + x) ^ 2 / x ^ 2) ≤ x + h : ((x + x) ^ (2 : ℝ) / x ^ (2 : ℝ)) ≤ x time_cmd equivalence divPowRevConstrRed/divPowRevConstrAuto : divPowRevConstr := by convexify @@ -967,7 +967,7 @@ time_cmd equivalence divPowRevConstrRed/divPowRevConstrAuto : divPowRevConstr := -- pow_sub-rev (obj) def powSubRevObj := optimization (x : ℝ) - minimize ((x ^ 2) / (sqrt x) : ℝ) + minimize ((x ^ (2 : ℝ)) / (sqrt x) : ℝ) subject to hx : 0 < x @@ -982,7 +982,7 @@ def powSubRevConstr := minimize (0 : ℝ) subject to hx : 0 < x - h : ((x ^ 2) / (sqrt x)) ≤ 1 + h : ((x ^ (2 : ℝ)) / (sqrt x)) ≤ 1 time_cmd equivalence powSubRevConstrRed/powSubRevConstrAuto : powSubRevConstr := by convexify @@ -992,7 +992,7 @@ time_cmd equivalence powSubRevConstrRed/powSubRevConstrAuto : powSubRevConstr := -- div_pow_eq_mul_pow_neg (obj) def divPowEqMulPowNegObj := optimization (x : ℝ) - minimize (1 / (x ^ 2) : ℝ) + minimize (1 / (x ^ (2 : ℝ)) : ℝ) subject to hx : 0 < x @@ -1007,7 +1007,7 @@ def divPowEqMulPowNegConstr := minimize (0 : ℝ) subject to hx : 0 < x - h : 1 / (x ^ 2) ≤ 1 + h : 1 / (x ^ (2 : ℝ)) ≤ 1 time_cmd equivalence divPowEqMulPowNegConstrRed/divPowEqMulPowNegConstrAuto : divPowEqMulPowNegConstr := by convexify @@ -1067,7 +1067,7 @@ time_cmd equivalence sqrtEqRpowConstrRed/sqrtEqRpowConstrAuto : sqrtEqRpowConstr -- pow_half_two (obj) def powHalfTwoObj := optimization (x : ℝ) - minimize ((sqrt x) ^ 2 : ℝ) + minimize ((sqrt x) ^ (2 : ℝ) : ℝ) subject to hx : 0 < x @@ -1082,7 +1082,7 @@ def powHalfTwoConstr := minimize (0 : ℝ) subject to hx : 0 < x - h : (sqrt x) ^ 2 ≤ 1 + h : (sqrt x) ^ (2 : ℝ) ≤ 1 time_cmd equivalence powHalfTwoConstrRed/powHalfTwoConstrAuto : powHalfTwoConstr := by convexify @@ -1220,7 +1220,7 @@ time_cmd equivalence expMulConstrRed/expMulConstrAuto : expMulConstr := by -- exp_mul-rev (obj) def expMulRevObj := optimization (x : ℝ) - minimize ((exp x) ^ 2 : ℝ) + minimize ((exp x) ^ (2 : ℝ) : ℝ) subject to hx : 0 < x @@ -1234,7 +1234,7 @@ def expMulRevConstr := optimization (x : ℝ) minimize (0 : ℝ) subject to - h : ((exp x) ^ 2) ≤ 1 + h : ((exp x) ^ (2 : ℝ)) ≤ 1 time_cmd equivalence expMulRevConstrRed/expMulRevConstrAuto : expMulRevConstr := by convexify @@ -1368,7 +1368,7 @@ time_cmd equivalence logDivConstrRed/logDivConstrAuto : logDivConstr := by -- log_div-rev (obj) def logDivRevObj := optimization (x : ℝ) - minimize (- (log (x ^ 2) - log x)) + minimize (- (log (x ^ (2 : ℝ)) - log x)) subject to hx : 0 < x @@ -1383,7 +1383,7 @@ def logDivRevConstr := minimize (0 : ℝ) subject to hx : 0 < x - h : - (log (x ^ 2) - log x) ≤ 1 + h : - (log (x ^ (2 : ℝ)) - log x) ≤ 1 time_cmd equivalence logDivRevConstrRed/logDivRevConstrAuto : logDivRevConstr := by convexify diff --git a/CvxLean/Test/DCP/Dcp.lean b/CvxLean/Test/DCP/Dcp.lean index df74a178..faaedff5 100644 --- a/CvxLean/Test/DCP/Dcp.lean +++ b/CvxLean/Test/DCP/Dcp.lean @@ -192,7 +192,7 @@ noncomputable def testEntr : Solution $ noncomputable def testQuadOverLin : Solution $ optimization (x y : ℝ) - minimize x ^ 2 / y + minimize x ^ (2 : ℝ) / y subject to c0 : 0 ≤ x c0 : 1 / 100000 ≤ y From 35e6bb8cd1c941a1050b33c81d0e6496d8c9ac45 Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 15 Dec 2023 16:16:48 -0500 Subject: [PATCH 188/189] feat: random names for solver files to avoid race conditions --- CvxLean/Tactic/Solver/Conic.lean | 76 ++++++++++++++++++-------------- 1 file changed, 43 insertions(+), 33 deletions(-) diff --git a/CvxLean/Tactic/Solver/Conic.lean b/CvxLean/Tactic/Solver/Conic.lean index 21d37417..fd093981 100644 --- a/CvxLean/Tactic/Solver/Conic.lean +++ b/CvxLean/Tactic/Solver/Conic.lean @@ -101,39 +101,49 @@ unsafe def conicSolverFromValues (goalExprs : SolutionExpr) cbf := cbf.setScalarConstraints (CBF.ConeProduct.mk n groupedCones.length groupedCones) - -- TODO: locking? - let outputPath := "solver/problem.sol" - IO.FS.withFile outputPath IO.FS.Mode.readWrite fun handle => do - -- Write input. - let inputPath := "solver/problem.cbf" - IO.FS.writeFile inputPath (ToString.toString cbf) - - -- Adjust path to MOSEK. - let p := if let some p' := ← IO.getEnv "PATH" then - if mosekBinPath != "" then p' ++ ":" ++ mosekBinPath else p' - else - mosekBinPath - - -- Run solver. - let out ← IO.Process.output { - cmd := "mosek", - args := #[inputPath], - env := #[("PATH", p)] } - if out.exitCode != 0 then - dbg_trace ("MOSEK exited with code " ++ ToString.toString out.exitCode) - return Sol.Response.failure out.exitCode.toNat - - let res := out.stdout - IO.println res - - -- Read output. - let output ← IO.FS.Handle.readToEnd handle - - match Sol.Parser.parse output with - | Except.ok res => return Sol.Response.success res - | Except.error err => - dbg_trace ("MOSEK output parsing failed. " ++ err) - return Sol.Response.failure 1 + let r ← IO.rand 0 (2 ^ 32 - 1) + let outputPath := s!"solver/problem{r}.sol" + let inputPath := s!"solver/problem{r}.cbf" + IO.FS.writeFile inputPath "" + IO.FS.writeFile outputPath "" + IO.FS.withFile outputPath IO.FS.Mode.read fun outHandle => do + IO.FS.withFile inputPath IO.FS.Mode.write fun inHandle => do + -- Write input. + inHandle.putStr (ToString.toString cbf) + + -- Adjust path to MOSEK. + let p := if let some p' := ← IO.getEnv "PATH" then + if mosekBinPath != "" then p' ++ ":" ++ mosekBinPath else p' + else + mosekBinPath + + -- Run solver. + let out ← IO.Process.output { + cmd := "mosek", + args := #[inputPath], + env := #[("PATH", p)] } + + dbg_trace s!"OUT: {out.stdout} {out.stderr}" + + if out.exitCode != 0 then + dbg_trace ("MOSEK exited with code " ++ ToString.toString out.exitCode) + return Sol.Response.failure out.exitCode.toNat + + let res := out.stdout + IO.println res + + -- Read output. + let output ← outHandle.readToEnd + + -- Remove temporary files. + IO.FS.removeFile inputPath + IO.FS.removeFile outputPath + + match Sol.Parser.parse output with + | Except.ok res => return Sol.Response.success res + | Except.error err => + dbg_trace ("MOSEK output parsing failed. " ++ err) + return Sol.Response.failure 1 /-- TODO: Move to Generation? -/ unsafe def exprFromSol (goalExprs : SolutionExpr) (sol : Sol.Result) : MetaM Expr := do From b829b518743a00e7616bc6a0487efa6ff8bf4b0a Mon Sep 17 00:00:00 2001 From: Ramon Fernandez Mir Date: Fri, 15 Dec 2023 16:32:17 -0500 Subject: [PATCH 189/189] fix: SO test power type --- CvxLean/Tactic/Solver/Conic.lean | 2 -- CvxLean/Test/Problems/SO.lean | 2 +- 2 files changed, 1 insertion(+), 3 deletions(-) diff --git a/CvxLean/Tactic/Solver/Conic.lean b/CvxLean/Tactic/Solver/Conic.lean index fd093981..2e45b1b8 100644 --- a/CvxLean/Tactic/Solver/Conic.lean +++ b/CvxLean/Tactic/Solver/Conic.lean @@ -123,8 +123,6 @@ unsafe def conicSolverFromValues (goalExprs : SolutionExpr) args := #[inputPath], env := #[("PATH", p)] } - dbg_trace s!"OUT: {out.stdout} {out.stderr}" - if out.exitCode != 0 then dbg_trace ("MOSEK exited with code " ++ ToString.toString out.exitCode) return Sol.Response.failure out.exitCode.toNat diff --git a/CvxLean/Test/Problems/SO.lean b/CvxLean/Test/Problems/SO.lean index 9c74b3aa..12ce635b 100644 --- a/CvxLean/Test/Problems/SO.lean +++ b/CvxLean/Test/Problems/SO.lean @@ -9,7 +9,7 @@ noncomputable def so1 := maximize sqrt (x - y) subject to c1 : y = 2 * x - 3 - c2 : x ^ 2 ≤ 2 + c2 : x ^ (2 : ℝ) ≤ 2 c3 : 0 ≤ x - y solve so1