Skip to content

Commit a051ac5

Browse files
robertylewisdigama0
andcommitted
feat: improve polyrith by testing for membership in the radical (#7790)
In @hrmacbeth 's tutorial on `polyrith`, there were examples of problems that it could almost solve, but failed. The goal was not expressible as a linear combination of the hypotheses but a power of the goal was. ```lean example (x y z : ℚ) (h : x = y) (h2 : x * y = 0) : x + y*z = 0 := sorry ``` Mathematically, `x+y*z` is in the radical of the ideal generated by `x-y, x*y`. There's a "standard trick" for testing membership in the radical without a search for the proper exponent: see e.g. section 2.2 of [arxiv.org/pdf/1007.3615.pdf](https://arxiv.org/pdf/1007.3615.pdf) or 4.2 Prop 8 of Cox, Little, O'Shea. This PR implements the trick in the Sage call made by `polyrith`. When the power returned is `n > 1`, we use `linear_combination (exp := n)` to check the certificate (#7789 ). The `polyrith` test infrastructure still needs to be ported from mathlib3. All tests in the test file succeed when they are uncommented. A future PR will restore the old test suite. Co-authored-by: Rob Lewis <rob.y.lewis@gmail.com> Co-authored-by: Mario Carneiro <di.gama@gmail.com>
1 parent 9183ab2 commit a051ac5

File tree

3 files changed

+63
-12
lines changed

3 files changed

+63
-12
lines changed

Mathlib/Tactic/Polyrith.lean

Lines changed: 20 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -33,6 +33,11 @@ It then calls a Python file that uses the SageMath API to compute the coefficien
3333
coefficients are then sent back to Lean, which parses them into pexprs. The information is then
3434
given to the `linear_combination` tactic, which completes the process by checking the certificate.
3535
36+
In fact, `polyrith` uses Sage to test for membership in the *radical* of the ideal.
37+
This means it searches for a linear combination of hypotheses that add up to a *power* of the goal.
38+
When this power is not 1, it uses the `(exp := n)` feature of `linear_combination` to report the
39+
certificate.
40+
3641
`polyrith` calls an external python script `scripts/polyrith_sage.py`. Because this is not a Lean
3742
file, changes to this script may not be noticed during Lean compilation if you have already
3843
generated olean files. If you are modifying this python script, you likely know what you're doing;
@@ -250,14 +255,24 @@ instance : FromJson Poly where
250255
mon := mon.mul' (.pow' (← fromJson? (← j.getArrVal? 0)) (← fromJson? (← j.getArrVal? 1)))
251256
pure mon
252257

258+
/-- A schema for the data reported by the Sage calculation -/
259+
structure SageCoeffAndPower where
260+
/-- The function call produces an array of polynomials
261+
parallel to the input list of hypotheses. -/
262+
coeffs : Array Poly
263+
/-- Sage produces an exponent (default 1) in the case where the hypothesess
264+
sum to a power of the goal. -/
265+
power : ℕ
266+
deriving FromJson, Repr
267+
253268
/-- The result of a sage call in the success case. -/
254269
structure SageSuccess where
255270
/-- The script returns a string containing python script to be sent to the remote server,
256271
when the tracing option is set. -/
257272
trace : Option String := none
258273
/-- The main result of the function call is an array of polynomials
259-
parallel to the input list of hypotheses. -/
260-
data : Option (Array Poly) := none
274+
parallel to the input list of hypotheses and an exponent for the goal. -/
275+
data : Option SageCoeffAndPower := none
261276
deriving FromJson, Repr
262277

263278
/-- The result of a sage call in the failure case. -/
@@ -338,7 +353,7 @@ def polyrith (g : MVarId) (only : Bool) (hyps : Array Expr)
338353
match ← sageOutput (createSageArgs traceOnly α vars hyps' tgt) with
339354
| .ok { trace, data } =>
340355
if let some trace := trace then logInfo trace
341-
if let some polys := data then
356+
if let some {coeffs := polys, power := pow} := data then
342357
let vars ← liftM <| (← get).atoms.mapM delab
343358
let p ← Poly.sumM (polys.zip hyps') fun (p, src, _) => do
344359
let h := .hyp (← delab (match src with | .input i => hyps[i]! | .fvar h => .fvar h))
@@ -348,7 +363,8 @@ def polyrith (g : MVarId) (only : Bool) (hyps : Array Expr)
348363
let stx := (withRef (← getRef) <| p.toSyntax vars).run
349364
let tac ←
350365
if let .const 0 := p then `(tactic| linear_combination)
351-
else `(tactic| linear_combination $stx:term)
366+
else if pow = 1 then `(tactic| linear_combination $stx:term)
367+
else `(tactic| linear_combination (exp := $(quote pow)) $stx:term)
352368
try
353369
guard (← Elab.runTactic g tac).1.isEmpty
354370
catch _ => throwError

scripts/polyrith_sage.py

Lines changed: 23 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -19,19 +19,29 @@ def create_query(type: str, n_vars: int, eq_list, goal_type):
1919
""" Create a query to invoke Sage's `MPolynomial_libsingular.lift`. See
2020
https://github.com/sagemath/sage/blob/f8df80820dc7321dc9b18c9644c3b8315999670b/src/sage/rings/polynomial/multi_polynomial_libsingular.pyx#L4472-L4518
2121
for a description of this method. """
22-
var_list = ", ".join([f"var{i}" for i in range(n_vars)])
22+
var_list = [f"var{i}" for i in range(n_vars)] + ['aux']
2323
query = f'''
2424
if {n_vars!r} != 0:
25-
P = PolynomialRing({type_str(type)}, 'var', {n_vars!r})
26-
[{var_list}] = P.gens()
25+
P = PolynomialRing({type_str(type)}, {var_list})
26+
[{", ".join(var_list)}] = P.gens()
27+
p = P({goal_type})
28+
gens = {eq_list} + [1 - p*aux]
29+
I = P.ideal(gens)
30+
coeffs = P(1).lift(I)
31+
power = max(cf.degree(aux) for cf in coeffs)
32+
coeffs = [P(cf.subs(aux = 1/p)*p^power) for cf in coeffs[:int(-1)]]
33+
print(str(power)+';'+serialize_polynomials(coeffs))
2734
else:
2835
# workaround for a Sage shortcoming with `n_vars = 0`,
2936
# `TypeError: no conversion of this ring to a Singular ring defined`
37+
# In this case, there is no need to look for membership in the *radical*;
38+
# we just check for membership in the ideal, and return exponent 1
39+
# if coefficients are found.
3040
P = PolynomialRing({type_str(type)}, 'var', 1)
31-
p = P({goal_type})
32-
I = P.ideal({eq_list})
33-
coeffs = p.lift(I)
34-
print(serialize_polynomials(coeffs))
41+
p = P({goal_type})
42+
I = P.ideal({eq_list})
43+
coeffs = p.lift(I)
44+
print('1;'+serialize_polynomials(coeffs))
3545
'''
3646
return query
3747

@@ -42,12 +52,17 @@ def __init__(self, ename, evalue, message='Error in Sage communication'):
4252
self.message = message
4353
super().__init__(self.message)
4454

55+
def parse_response(resp: str) -> str:
56+
exp, data = resp.split(';', 1)
57+
return dict(power=int(exp), coeffs=json.loads(data))
58+
59+
4560
def evaluate_in_sage(query: str) -> str:
4661
data = {'code': query}
4762
headers = {'content-type': 'application/x-www-form-urlencoded'}
4863
response = requests.post('https://sagecell.sagemath.org/service', data, headers=headers).json()
4964
if response['success']:
50-
return json.loads(response.get('stdout'))
65+
return parse_response(response.get('stdout'))
5166
elif 'execute_reply' in response and 'ename' in response['execute_reply'] and 'evalue' in response['execute_reply']:
5267
raise EvaluationError(response['execute_reply']['ename'], response['execute_reply']['evalue'])
5368
else:

test/polyrith.lean

Lines changed: 20 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -561,6 +561,26 @@ by polyrith
561561
-- polyrith,
562562
-- end
563563
564+
/-
565+
566+
### Examples with exponent
567+
568+
-/
569+
570+
example (x y z : ℚ) (h : x = y) (h2 : x * y = 0) : x + y*z = 0 := by
571+
polyrith
572+
573+
example (K : Type)
574+
[Field K]
575+
[CharZero K]
576+
{x y z : K}
577+
(h₂ : y ^ 3 + x * (3 * z ^ 2) = 0)
578+
(h₁ : x ^ 3 + z * (3 * y ^ 2) = 0)
579+
(h₀ : y * (3 * x ^ 2) + z ^ 3 = 0)
580+
(h : x ^ 3 * y + y ^ 3 * z + z ^ 3 * x = 0) :
581+
x = 0 := by
582+
polyrith
583+
564584
/-!
565585
### With trace enabled
566586
Here, the tactic will trace the command that gets sent to sage,

0 commit comments

Comments
 (0)