Skip to content

Commit

Permalink
fix(tactic/polyrith): fix crash when hypotheses are constant polynomials
Browse files Browse the repository at this point in the history
  • Loading branch information
eric-wieser committed Oct 9, 2023
1 parent 159651e commit 62823a4
Show file tree
Hide file tree
Showing 3 changed files with 27 additions and 8 deletions.
2 changes: 0 additions & 2 deletions Mathlib/Tactic/Polyrith.lean
Original file line number Diff line number Diff line change
Expand Up @@ -332,8 +332,6 @@ def polyrith (g : MVarId) (only : Bool) (hyps : Array Expr)
if hyps'.isEmpty then
return ← byRing "polyrith did not find any relevant hypotheses"
let vars := (← get).atoms.size
if vars = 0 then
return ← byRing "polyrith did not find find any variables"
match ← sageOutput (createSageArgs traceOnly α vars hyps' tgt) with
| .ok { trace, data } =>
if let some trace := trace then logInfo trace
Expand Down
11 changes: 8 additions & 3 deletions scripts/polyrith_sage.py
Original file line number Diff line number Diff line change
Expand Up @@ -21,9 +21,14 @@ def create_query(type: str, n_vars: int, eq_list, goal_type):
for a description of this method. """
var_list = ", ".join([f"var{i}" for i in range(n_vars)])
query = f'''
P = PolynomialRing({type_str(type)}, 'var', {n_vars!r})
[{var_list}] = P.gens()
gens = {eq_list}
if {n_vars!r} != 0:
P = PolynomialRing({type_str(type)}, 'var', {n_vars!r})
[{var_list}] = P.gens()
else:
# workaround for a Sage bug
P = PolynomialRing({type_str(type)}, 'var', 1)
# ensure that the equalities are cast to polynomials
gens = [P() + eq for eq in {eq_list}]
p = P({goal_type})
I = ideal(gens)
coeffs = p.lift(I)
Expand Down
22 changes: 19 additions & 3 deletions test/polyrith.lean
Original file line number Diff line number Diff line change
Expand Up @@ -433,11 +433,27 @@ A full test suite is provided at the bottom of the file.
-- "(1 - 2)"]
-- "linear_combination h1 - h2"

-- example {R} [CommRing R] (x : R) (h2 : (2 : R) = 0) : x + x = 0 :=
-- by test_polyrith
-- "{\"data\":[\"(poly.var 0)\"],\"success\":true}"
-- ["ff",
-- "R",
-- "1",
-- "[(2 - 0)]",
-- "((var0 + var0) - 0)"]
-- "linear_combination x * h2"

-- We comment the following tests so that we don't overwhelm the SageCell API.


-- example {R} [CommRing R] (_x : R) (h : (2 : R) = 4) : (0 : R) = 2 :=
-- by test_polyrith
-- "{\"data\":[\"(poly.const 1/1)\"],\"success\":true}"
-- ["ff",
-- "R",
-- "0",
-- "[(2 - 4)]",
-- "(0 - 2)"]
-- "linear_combination h"

-- We comment the following tests so that we don't overwhelm the SageCell API.

/-
Expand Down

0 comments on commit 62823a4

Please sign in to comment.