Skip to content

Commit 2f32110

Browse files
authored
feat: support for Rat scientific literals (#10961)
This PR adds support for scientific literals for `Rat` in `grind`. `grind` does not yet add support for this kind of literal in arbitrary fields. closes #10489
1 parent cdaa827 commit 2f32110

File tree

7 files changed

+105
-12
lines changed

7 files changed

+105
-12
lines changed

src/Init/Data/Rat/Lemmas.lean

Lines changed: 7 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1011,7 +1011,6 @@ theorem intCast_neg_iff {a : Int} :
10111011
/--
10121012
Alternative statement of `ofScientific_def`.
10131013
-/
1014-
@[grind =]
10151014
theorem ofScientific_def' :
10161015
(OfScientific.ofScientific m s e : Rat) = m * (10 ^ (if s then -e else e : Int)) := by
10171016
change Rat.ofScientific _ _ _ = _
@@ -1023,6 +1022,13 @@ theorem ofScientific_def' :
10231022
· push_cast
10241023
rfl
10251024

1025+
theorem ofScientific_def_eq_if :
1026+
(OfScientific.ofScientific m s e : Rat) = if s then (m : Rat) / (10 : Rat) ^ e else (m : Rat) * (10 : Rat) ^ e := by
1027+
simp [ofScientific_def']
1028+
split
1029+
next => rw [Rat.zpow_neg, ← Rat.div_def, Rat.zpow_natCast]
1030+
next => rw [Rat.zpow_natCast]
1031+
10261032
/-!
10271033
# min and max
10281034
-/

src/Init/Grind/Norm.lean

Lines changed: 4 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -4,11 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Leonardo de Moura
55
-/
66
module
7-
87
prelude
98
public import Init.Data.Int.Linear
109
public import Init.Grind.Ring.Field
11-
10+
public import Init.Data.Rat.Lemmas
1211
public section
1312

1413
namespace Lean.Grind
@@ -153,7 +152,7 @@ init_grind_norm
153152
/- Pre theorems -/
154153
|
155154
/- Post theorems -/
156-
iff_eq heq_eq_eq
155+
iff_eq heq_eq_eq eq_self
157156
-- And
158157
and_true true_and and_false false_and and_assoc
159158
-- ite
@@ -208,5 +207,7 @@ init_grind_norm
208207
Ring.intCast_mul
209208
Ring.intCast_pow
210209
Ring.intCast_sub
210+
-- Rationals
211+
Rat.ofScientific_def_eq_if Rat.zpow_neg
211212

212213
end Lean.Grind

src/Lean/Meta/LitValues.lean

Lines changed: 20 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -4,14 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Leonardo de Moura
55
-/
66
module
7-
87
prelude
98
public import Lean.Meta.Basic
10-
9+
public import Init.Data.Rat.Basic
1110
public section
12-
1311
namespace Lean.Meta
14-
1512
/-!
1613
Helper functions for recognizing builtin literal values.
1714
This module focus on recognizing the standard representation used in Lean for these literals.
@@ -53,6 +50,25 @@ def getIntValue? (e : Expr) : MetaM (Option Int) := do
5350
let some (n, _) ← getOfNatValue? a ``Int | return none
5451
return some (-↑n)
5552

53+
/--
54+
Return `some i` if `e` `OfNat.ofNat`-application encoding a rational, or `Neg.neg`-application of one,
55+
or a division.
56+
-/
57+
def getRatValue? (e : Expr) : MetaM (Option Rat) := do
58+
match_expr e with
59+
| HDiv.hDiv _ _ _ _ a b =>
60+
let some n ← getRatValueNum? a | return none
61+
let some (d, _) ← getOfNatValue? b ``Rat | return none
62+
return some (n / (d : Rat))
63+
| _ => getRatValueNum? e
64+
where
65+
getRatValueNum? (e : Expr) : MetaM (Option Rat) := do
66+
if let some (n, _) ← getOfNatValue? e ``Rat then
67+
return some (n : Rat)
68+
let_expr Neg.neg _ _ a ← e | return none
69+
let some (n, _) ← getOfNatValue? a ``Rat | return none
70+
return some (- (n : Rat))
71+
5672
/-- Return `some c` if `e` is a `Char.ofNat`-application that encodes the character `c`. -/
5773
def getCharValue? (e : Expr) : MetaM (Option Char) := do
5874
let_expr Char.ofNat n ← e | return none

src/Lean/Meta/Tactic/Grind/Arith/Simproc.lean

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -168,6 +168,18 @@ builtin_simproc_decl normIntCastNum (IntCast.intCast _) := fun e => do
168168
let h := mkApp4 (mkConst ``Grind.Ring.intCast_eq_ofNat_of_nonneg us) α ringInst a eagerReflBoolTrue
169169
return .done { expr := n, proof? := some h }
170170

171+
builtin_dsimproc [simp, seval] normPowRatInt ((_ : Rat) ^ (_ : Int)) := fun e => do
172+
let_expr HPow.hPow _ _ _ _ a b ← e | return .continue
173+
let some v₁ ← getRatValue? a | return .continue
174+
let some v₂ ← getIntValue? b | return .continue
175+
let warning := (← Simp.getConfig).warnExponents
176+
unless (← checkExponent v₂.natAbs (warning := warning)) do return .continue
177+
if v₂ < 0 then
178+
-- **Note**: we use `Rat.zpow_neg` as a normalization rule
179+
return .continue
180+
else
181+
return .done <| toExpr (v₁ ^ v₂)
182+
171183
/-!
172184
Add additional arithmetic simprocs
173185
-/
@@ -193,6 +205,7 @@ def addSimproc (s : Simprocs) : CoreM Simprocs := do
193205
let s ← s.add ``normIntOfNatInst (post := false)
194206
let s ← s.add ``normNatCastNum (post := false)
195207
let s ← s.add ``normIntCastNum (post := false)
208+
let s ← s.add ``normPowRatInt (post := false)
196209
return s
197210

198211
end Lean.Meta.Grind.Arith

src/Lean/ToExpr.lean

Lines changed: 21 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -4,13 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Leonardo de Moura
55
-/
66
module
7-
87
prelude
98
public import Lean.ToLevel
10-
9+
public import Init.Data.Rat.Basic
1110
public section
1211
universe u
13-
1412
namespace Lean
1513

1614
/--
@@ -51,6 +49,26 @@ where
5149
mkApp3 (.const ``OfNat.ofNat [0]) (.const ``Int []) r
5250
(.app (.const ``instOfNat []) r)
5351

52+
instance : ToExpr Rat where
53+
toTypeExpr := .const ``Rat []
54+
toExpr i := if i.den == 1 then
55+
mkInt i.num
56+
else
57+
mkApp6 (.const ``HDiv.hDiv [0, 0, 0]) (.const ``Rat []) (.const ``Rat []) (.const ``Rat [])
58+
(mkApp2 (.const ``instHDiv [0]) (.const ``Rat []) (.const ``Rat.instDiv []))
59+
(mkInt i.num) (mkInt i.den)
60+
where
61+
mkInt (i : Int) : Expr :=
62+
if 0 ≤ i then
63+
mkNat i.toNat
64+
else
65+
mkApp3 (.const ``Neg.neg [0]) (.const ``Rat []) (.const ``Rat.instNeg [])
66+
(mkNat (-i).toNat)
67+
mkNat (n : Nat) : Expr :=
68+
let r := mkRawNatLit n
69+
mkApp3 (.const ``OfNat.ofNat [0]) (.const ``Rat []) r
70+
(.app (.const ``Rat.instOfNat []) r)
71+
5472
instance : ToExpr (Fin n) where
5573
toTypeExpr := .app (mkConst ``Fin) (toExpr n)
5674
toExpr a :=

tests/lean/run/grind_10489.lean

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,10 @@
1+
example : (2 / 3 : Rat) ≤ (0.67 : Rat) := by grind
2+
example : (1.2 : Rat) ≤ (1.21 : Rat) := by grind
3+
example : (2 / 3 : Rat) ≤ (67 / 100 : Rat) := by grind
4+
example : (2 / 3 : Rat) ≤ (67 * 10 ^ (-2) : Rat) := by grind
5+
example : (2 / 3 : Rat) ≤ (67 / 10 ^ 2 : Rat) := by grind
6+
example : (2 / 3 : Rat) ≤ (67 / 10 ^ (2:Int) : Rat) := by grind
7+
example : (1.2345 : Rat) ≤ (1.2346 : Rat) := by grind
8+
example : (1.2345 : Rat) ≤ (1.234501 : Rat) := by grind
9+
example : (2.3 : Rat) ≤ (4.5 : Rat) := by grind
10+
example : (2.3 : Rat) ≤ (5/2 : Rat) := by grind

tests/lean/run/toExpr.lean

Lines changed: 30 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -3,7 +3,6 @@ import Lean
33
open Lean
44

55
unsafe def test {α : Type} [ToString α] [ToExpr α] [BEq α] (a : α) : CoreM Unit := do
6-
let env ← getEnv;
76
let auxName := `_toExpr._test;
87
let decl := Declaration.defnDecl {
98
name := auxName,
@@ -30,3 +29,33 @@ match newEnv.evalConst α {} auxName with
3029
#eval test ['a', 'b', 'c']
3130
#eval test ("hello", true)
3231
#eval test ((), 10)
32+
#eval test (1:Rat)
33+
#eval test (-1:Rat)
34+
#eval test (2:Rat)
35+
#eval test (-2:Rat)
36+
#eval test (1/(-2):Rat)
37+
#eval test (-2/3:Rat)
38+
#eval test (-2/1:Rat)
39+
#eval test (-20/3:Rat)
40+
#eval test (-1.234:Rat)
41+
#eval test (0.67:Rat)
42+
43+
open Lean Meta
44+
def testRat (n : Rat) : MetaM Unit := do
45+
let e := toExpr n
46+
IO.println e
47+
let some n' ← getRatValue? e | throwError "`Rat` expected{indentExpr e}"
48+
IO.println n'
49+
assert! n == n'
50+
51+
#eval testRat 0
52+
#eval testRat 1
53+
#eval testRat (1/2)
54+
#eval testRat (1/(-2))
55+
#eval testRat (2/3)
56+
#eval testRat (0.67)
57+
#eval testRat (1.67)
58+
#eval testRat (1.68)
59+
#eval testRat (-1.67)
60+
#eval testRat (-2)
61+
#eval testRat (-0.67)

0 commit comments

Comments
 (0)