File tree Expand file tree Collapse file tree 3 files changed +34
-1
lines changed
Expand file tree Collapse file tree 3 files changed +34
-1
lines changed Original file line number Diff line number Diff line change 77prelude
88public import Init.Grind.Ring.Basic
99public import Init.Grind.Ring.Field
10+ public import Init.Grind.Ring.OfScientific
1011public import Init.Grind.Ring.Envelope
1112public import Init.Grind.Ring.CommSolver
1213public import Init.Grind.Ring.CommSemiringAdapter
Original file line number Diff line number Diff line change 1+ /-
2+ Copyright (c) 2025 Lean FRO, LLC. or its affiliates. All Rights Reserved.
3+ Released under Apache 2.0 license as described in the file LICENSE.
4+ Authors: Kim Morrison
5+ -/
6+ module
7+
8+ prelude
9+ public import Init.Grind.Ring.Field
10+ public import Init.Data.OfScientific
11+
12+ public section
13+
14+ namespace Lean.Grind
15+
16+ attribute [local instance] Semiring.natCast
17+
18+ /--
19+ A type class that ensures that `OfScientific.ofScientific` is properly defined with respect to
20+ the field structure.
21+ -/
22+ class LawfulOfScientific (α : Type u) [Field α] [OfScientific α] : Prop where
23+ /--
24+ `OfScientific.ofScientific` is properly defined with respect to the field structure.
25+ -/
26+ ofScientific_def :
27+ OfScientific.ofScientific m s e = if s then (Nat.cast m : α) / 10 ^e else (Nat.cast m : α) * 10 ^e
28+
29+ end Lean.Grind
Original file line number Diff line number Diff line change @@ -6,7 +6,7 @@ Authors: Robin Arnez
66module
77
88prelude
9- public import Init.Grind.Ring.Field
9+ public import Init.Grind.Ring.OfScientific
1010public import Init.Data.Rat.Lemmas
1111
1212public section
@@ -56,4 +56,7 @@ instance : NoNatZeroDivisors Rat where
5656 change k * a = k * b at h₂
5757 simpa [← Rat.mul_assoc, Rat.inv_mul_cancel, h₁] using congrArg ((k : Rat)⁻¹ * ·) h₂
5858
59+ instance : LawfulOfScientific Rat where
60+ ofScientific_def {m s e} := by rw [Rat.ofScientific_def_eq_if]
61+
5962end Lean.Grind
You can’t perform that action at this time.
0 commit comments