Skip to content

Commit 87cc330

Browse files
authored
feat: ordered ring typeclass for grind (#8429)
This PR adds `Lean.Grind.Ring.IsOrdered`, and cleans up the ring/module grind API. These typeclasses are at present unused, but will support future algorithmic improvements in `grind`.
1 parent 47a1355 commit 87cc330

File tree

9 files changed

+279
-28
lines changed

9 files changed

+279
-28
lines changed

src/Init/Grind.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -16,4 +16,5 @@ import Init.Grind.Offset
1616
import Init.Grind.PP
1717
import Init.Grind.CommRing
1818
import Init.Grind.Module
19+
import Init.Grind.Ordered
1920
import Init.Grind.Ext

src/Init/Grind/CommRing/Basic.lean

Lines changed: 6 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -104,6 +104,12 @@ theorem ofNat_mul (a b : Nat) : OfNat.ofNat (α := α) (a * b) = OfNat.ofNat a *
104104
theorem natCast_mul (a b : Nat) : ((a * b : Nat) : α) = ((a : α) * (b : α)) := by
105105
rw [← ofNat_eq_natCast, ofNat_mul, ofNat_eq_natCast, ofNat_eq_natCast]
106106

107+
theorem pow_one (a : α) : a ^ 1 = a := by
108+
rw [pow_succ, pow_zero, one_mul]
109+
110+
theorem pow_two (a : α) : a ^ 2 = a * a := by
111+
rw [pow_succ, pow_one]
112+
107113
theorem pow_add (a : α) (k₁ k₂ : Nat) : a ^ (k₁ + k₂) = a^k₁ * a^k₂ := by
108114
induction k₂
109115
next => simp [pow_zero, mul_one]
@@ -274,7 +280,6 @@ instance : IntModule α where
274280
hmul_zero := by simp [mul_zero]
275281
hmul_add := by simp [left_distrib]
276282
mul_hmul := by simp [intCast_mul, mul_assoc]
277-
neg_hmul := by simp [intCast_neg, neg_mul]
278283
neg_add_cancel := by simp [neg_add_cancel]
279284
sub_eq_add_neg := by simp [sub_eq_add_neg]
280285

src/Init/Grind/Module.lean

Lines changed: 0 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -7,4 +7,3 @@ module
77

88
prelude
99
import Init.Grind.Module.Basic
10-
import Init.Grind.Module.Int

src/Init/Grind/Module/Basic.lean

Lines changed: 55 additions & 18 deletions
Original file line numberDiff line numberDiff line change
@@ -32,39 +32,76 @@ class IntModule (M : Type u) extends Zero M, Add M, Neg M, Sub M, HMul Int M M w
3232
zero_hmul : ∀ a : M, (0 : Int) * a = 0
3333
one_hmul : ∀ a : M, (1 : Int) * a = a
3434
add_hmul : ∀ n m : Int, ∀ a : M, (n + m) * a = n * a + m * a
35-
neg_hmul : ∀ n : Int, ∀ a : M, (-n) * a = - (n * a)
3635
hmul_zero : ∀ n : Int, n * (0 : M) = 0
3736
hmul_add : ∀ n : Int, ∀ a b : M, n * (a + b) = n * a + n * b
3837
mul_hmul : ∀ n m : Int, ∀ a : M, (n * m) * a = n * (m * a)
3938
neg_add_cancel : ∀ a : M, -a + a = 0
4039
sub_eq_add_neg : ∀ a b : M, a - b = a + -b
4140

41+
namespace IntModule
42+
4243
attribute [instance 100] IntModule.toZero IntModule.toAdd IntModule.toNeg IntModule.toSub IntModule.toHMul
4344

44-
instance IntModule.toNatModule (M : Type u) [i : IntModule M] : NatModule M :=
45+
instance toNatModule (M : Type u) [i : IntModule M] : NatModule M :=
4546
{ i with
4647
hMul a x := (a : Int) * x
4748
hmul_zero := by simp [IntModule.hmul_zero]
4849
add_hmul := by simp [IntModule.add_hmul]
4950
hmul_add := by simp [IntModule.hmul_add]
5051
mul_hmul := by simp [IntModule.mul_hmul] }
5152

52-
/-- A preorder is a reflexive, transitive relation `≤` with `a < b` defined in the obvious way. -/
53-
class Preorder (α : Type u) extends LE α, LT α where
54-
le_refl : ∀ a : α, a ≤ a
55-
le_trans : ∀ a b c : α, a ≤ b → b ≤ c → a ≤ c
56-
lt := fun a b => a ≤ b ∧ ¬b ≤ a
57-
lt_iff_le_not_le : ∀ a b : α, a < b ↔ a ≤ b ∧ ¬b ≤ a := by intros; rfl
58-
59-
class IntModule.IsOrdered (M : Type u) [Preorder M] [IntModule M] where
60-
neg_le_iff : ∀ a b : M, -a ≤ b ↔ -b ≤ a
61-
neg_lt_iff : ∀ a b : M, -a < b ↔ -b < a
62-
add_lt_left : ∀ a b c : M, a < b → a + c < b + c
63-
add_lt_right : ∀ a b c : M, a < b → c + a < c + b
64-
hmul_pos : ∀ (k : Int) (a : M), 0 < a → (0 < k ↔ 0 < k * a)
65-
hmul_neg : ∀ (k : Int) (a : M), a < 0 → (0 < k ↔ k * a < 0)
66-
hmul_nonneg : ∀ (k : Int) (a : M), 0 ≤ a → 0 ≤ k → 0 ≤ k * a
67-
hmul_nonpos : ∀ (k : Int) (a : M), a ≤ 00 ≤ k → k * a ≤ 0
53+
variable {M : Type u} [IntModule M]
54+
55+
theorem add_neg_cancel (a : M) : a + -a = 0 := by
56+
rw [add_comm, neg_add_cancel]
57+
58+
theorem add_left_inj {a b : M} (c : M) : a + c = b + c ↔ a = b :=
59+
fun h => by simpa [add_assoc, add_neg_cancel, add_zero] using (congrArg (· + -c) h),
60+
fun g => congrArg (· + c) g⟩
61+
62+
theorem add_right_inj (a b c : M) : a + b = a + c ↔ b = c := by
63+
rw [add_comm a b, add_comm a c, add_left_inj]
64+
65+
theorem neg_zero : (-0 : M) = 0 := by
66+
rw [← add_left_inj 0, neg_add_cancel, add_zero]
67+
68+
theorem neg_neg (a : M) : -(-a) = a := by
69+
rw [← add_left_inj (-a), neg_add_cancel, add_neg_cancel]
70+
71+
theorem neg_eq_zero (a : M) : -a = 0 ↔ a = 0 :=
72+
fun h => by
73+
replace h := congrArg (-·) h
74+
simpa [neg_neg, neg_zero] using h,
75+
fun h => by rw [h, neg_zero]⟩
76+
77+
theorem neg_add (a b : M) : -(a + b) = -a + -b := by
78+
rw [← add_left_inj (a + b), neg_add_cancel, add_assoc (-a), add_comm a b, ← add_assoc (-b),
79+
neg_add_cancel, zero_add, neg_add_cancel]
80+
81+
theorem neg_sub (a b : M) : -(a - b) = b - a := by
82+
rw [sub_eq_add_neg, neg_add, neg_neg, sub_eq_add_neg, add_comm]
83+
84+
theorem sub_self (a : M) : a - a = 0 := by
85+
rw [sub_eq_add_neg, add_neg_cancel]
86+
87+
theorem sub_eq_iff {a b c : M} : a - b = c ↔ a = c + b := by
88+
rw [sub_eq_add_neg]
89+
constructor
90+
next => intro; subst c; rw [add_assoc, neg_add_cancel, add_zero]
91+
next => intro; subst a; rw [add_assoc, add_comm b, neg_add_cancel, add_zero]
92+
93+
theorem sub_eq_zero_iff {a b : M} : a - b = 0 ↔ a = b := by
94+
simp [sub_eq_iff, zero_add]
95+
96+
theorem neg_hmul (n : Int) (a : M) : (-n) * a = - (n * a) := by
97+
apply (add_left_inj (n * a)).mp
98+
rw [← add_hmul, Int.add_left_neg, zero_hmul, neg_add_cancel]
99+
100+
theorem hmul_neg (n : Int) (a : M) : n * (-a) = - (n * a) := by
101+
apply (add_left_inj (n * a)).mp
102+
rw [← hmul_add, neg_add_cancel, neg_add_cancel, hmul_zero]
103+
104+
end IntModule
68105

69106
/--
70107
Special case of Mathlib's `NoZeroSMulDivisors Nat α`.

src/Init/Grind/Ordered.lean

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,12 @@
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+
import Init.Grind.Ordered.PartialOrder
10+
import Init.Grind.Ordered.Module
11+
import Init.Grind.Ordered.Ring
12+
import Init.Grind.Ordered.Int
Lines changed: 9 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -6,7 +6,7 @@ Authors: Kim Morrison
66
module
77

88
prelude
9-
import Init.Grind.Module.Basic
9+
import Init.Grind.Ordered.Ring
1010
import Init.Grind.CommRing.Int
1111
import Init.Omega
1212

@@ -18,17 +18,18 @@ namespace Lean.Grind
1818

1919
instance : Preorder Int where
2020
le_refl := Int.le_refl
21-
le_trans _ _ _ := Int.le_trans
21+
le_trans := Int.le_trans
2222
lt_iff_le_not_le := by omega
2323

2424
instance : IntModule.IsOrdered Int where
2525
neg_le_iff := by omega
26-
neg_lt_iff := by omega
27-
add_lt_left := by omega
28-
add_lt_right := by omega
26+
add_le_left := by omega
2927
hmul_pos k a ha := ⟨fun hk => Int.mul_pos hk ha, fun h => Int.pos_of_mul_pos_left h ha⟩
30-
hmul_neg k a ha := ⟨fun hk => Int.mul_neg_of_pos_of_neg hk ha, fun h => Int.pos_of_mul_neg_left h ha⟩
31-
hmul_nonpos k a ha hk := Int.mul_nonpos_of_nonneg_of_nonpos hk ha
32-
hmul_nonneg k a ha hk := Int.mul_nonneg hk ha
28+
hmul_nonneg hk ha := Int.mul_nonneg hk ha
29+
30+
instance : Ring.IsOrdered Int where
31+
zero_lt_one := by omega
32+
mul_lt_mul_of_pos_left := Int.mul_lt_mul_of_pos_left
33+
mul_lt_mul_of_pos_right := Int.mul_lt_mul_of_pos_right
3334

3435
end Lean.Grind

src/Init/Grind/Ordered/Module.lean

Lines changed: 69 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,69 @@
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+
import Init.Data.Int.Order
10+
import Init.Grind.Module.Basic
11+
import Init.Grind.Ordered.PartialOrder
12+
13+
namespace Lean.Grind
14+
15+
class IntModule.IsOrdered (M : Type u) [Preorder M] [IntModule M] where
16+
neg_le_iff : ∀ a b : M, -a ≤ b ↔ -b ≤ a
17+
add_le_left : ∀ {a b : M}, a ≤ b → (c : M) → a + c ≤ b + c
18+
hmul_pos : ∀ (k : Int) {a : M}, 0 < a → (0 < k ↔ 0 < k * a)
19+
hmul_nonneg : ∀ {k : Int} {a : M}, 0 ≤ k → 0 ≤ a → 0 ≤ k * a
20+
21+
namespace IntModule.IsOrdered
22+
23+
variable {M : Type u} [Preorder M] [IntModule M] [IntModule.IsOrdered M]
24+
25+
theorem le_neg_iff {a b : M} : a ≤ -b ↔ b ≤ -a := by
26+
conv => lhs; rw [← neg_neg a]
27+
rw [neg_le_iff, neg_neg]
28+
29+
theorem neg_lt_iff {a b : M} : -a < b ↔ -b < a := by
30+
simp [Preorder.lt_iff_le_not_le]
31+
rw [neg_le_iff, le_neg_iff]
32+
33+
theorem lt_neg_iff {a b : M} : a < -b ↔ b < -a := by
34+
conv => lhs; rw [← neg_neg a]
35+
rw [neg_lt_iff, neg_neg]
36+
37+
theorem neg_nonneg_iff {a : M} : 0 ≤ -a ↔ a ≤ 0 := by
38+
rw [le_neg_iff, neg_zero]
39+
40+
theorem neg_pos_iff {a : M} : 0 < -a ↔ a < 0 := by
41+
rw [lt_neg_iff, neg_zero]
42+
43+
theorem add_lt_left {a b : M} (h : a < b) (c : M) : a + c < b + c := by
44+
simp [Preorder.lt_iff_le_not_le] at h ⊢
45+
constructor
46+
· exact add_le_left h.1 _
47+
· intro w
48+
apply h.2
49+
replace w := add_le_left w (-c)
50+
rw [add_assoc, add_assoc, add_neg_cancel, add_zero, add_zero] at w
51+
exact w
52+
53+
theorem add_le_right (a : M) {b c : M} (h : b ≤ c) : a + b ≤ a + c := by
54+
rw [add_comm a b, add_comm a c]
55+
exact add_le_left h a
56+
57+
theorem add_lt_right (a : M) {b c : M} (h : b < c) : a + b < a + c := by
58+
rw [add_comm a b, add_comm a c]
59+
exact add_lt_left h a
60+
61+
theorem hmul_neg (k : Int) {a : M} (h : a < 0) : 0 < k ↔ k * a < 0 := by
62+
simpa [IntModule.hmul_neg, neg_pos_iff] using hmul_pos k (neg_pos_iff.mpr h)
63+
64+
theorem hmul_nonpos {k : Int} {a : M} (hk : 0 ≤ k) (ha : a ≤ 0) : k * a ≤ 0 := by
65+
simpa [IntModule.hmul_neg, neg_nonneg_iff] using hmul_nonneg hk (neg_nonneg_iff.mpr ha)
66+
67+
end IntModule.IsOrdered
68+
69+
end Lean.Grind
Lines changed: 61 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,61 @@
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+
import Init.Data.Int.Order
10+
11+
namespace Lean.Grind
12+
13+
/-- A preorder is a reflexive, transitive relation `≤` with `a < b` defined in the obvious way. -/
14+
class Preorder (α : Type u) extends LE α, LT α where
15+
le_refl : ∀ a : α, a ≤ a
16+
le_trans : ∀ {a b c : α}, a ≤ b → b ≤ c → a ≤ c
17+
lt := fun a b => a ≤ b ∧ ¬b ≤ a
18+
lt_iff_le_not_le : ∀ {a b : α}, a < b ↔ a ≤ b ∧ ¬b ≤ a := by intros; rfl
19+
20+
namespace Preorder
21+
22+
variable {α : Type u} [Preorder α]
23+
24+
theorem le_of_lt {a b : α} (h : a < b) : a ≤ b := (lt_iff_le_not_le.mp h).1
25+
26+
theorem lt_of_lt_of_le {a b c : α} (h₁ : a < b) (h₂ : b ≤ c) : a < c := by
27+
simp [lt_iff_le_not_le] at h₁ ⊢
28+
exact ⟨le_trans h₁.1 h₂, fun h => h₁.2 (le_trans h₂ h)⟩
29+
30+
theorem lt_of_le_of_lt {a b c : α} (h₁ : a ≤ b) (h₂ : b < c) : a < c := by
31+
simp [lt_iff_le_not_le] at h₂ ⊢
32+
exact ⟨le_trans h₁ h₂.1, fun h => h₂.2 (le_trans h h₁)⟩
33+
34+
theorem lt_trans {a b c : α} (h₁ : a < b) (h₂ : b < c) : a < c :=
35+
lt_of_lt_of_le h₁ (le_of_lt h₂)
36+
37+
theorem lt_irrefl {a : α} (h : a < a) : False := by
38+
simp [lt_iff_le_not_le] at h
39+
40+
end Preorder
41+
42+
class PartialOrder (α : Type u) extends Preorder α where
43+
le_antisymm : ∀ {a b : α}, a ≤ b → b ≤ a → a = b
44+
45+
namespace PartialOrder
46+
47+
variable {α : Type u} [PartialOrder α]
48+
49+
theorem le_iff_lt_or_eq {a b : α} : a ≤ b ↔ a < b ∨ a = b := by
50+
constructor
51+
· intro h
52+
rw [Preorder.lt_iff_le_not_le, Classical.or_iff_not_imp_right]
53+
exact fun w => ⟨h, fun w' => w (le_antisymm h w')⟩
54+
· intro h
55+
cases h with
56+
| inl h => exact Preorder.le_of_lt h
57+
| inr h => subst h; exact Preorder.le_refl a
58+
59+
end PartialOrder
60+
61+
end Lean.Grind

src/Init/Grind/Ordered/Ring.lean

Lines changed: 66 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,66 @@
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+
import Init.Grind.CommRing.Basic
10+
import Init.Grind.Ordered.Module
11+
12+
namespace Lean.Grind
13+
14+
class Ring.IsOrdered (R : Type u) [Ring R] [Preorder R] extends IntModule.IsOrdered R where
15+
/-- In a strict ordered semiring, we have `0 < 1`. -/
16+
zero_lt_one : 0 < 1
17+
/-- In a strict ordered semiring, we can multiply an inequality `a < b` on the left
18+
by a positive element `0 < c` to obtain `c * a < c * b`. -/
19+
mul_lt_mul_of_pos_left : ∀ {a b c : R}, a < b → 0 < c → c * a < c * b
20+
/-- In a strict ordered semiring, we can multiply an inequality `a < b` on the right
21+
by a positive element `0 < c` to obtain `a * c < b * c`. -/
22+
mul_lt_mul_of_pos_right : ∀ {a b c : R}, a < b → 0 < c → a * c < b * c
23+
24+
namespace Ring.IsOrdered
25+
26+
variable {R : Type u} [Ring R] [PartialOrder R] [Ring.IsOrdered R]
27+
28+
theorem mul_le_mul_of_nonneg_left {a b c : R} (h : a ≤ b) (h' : 0 ≤ c) : c * a ≤ c * b := by
29+
rw [PartialOrder.le_iff_lt_or_eq] at h'
30+
cases h' with
31+
| inl h' =>
32+
have p := mul_lt_mul_of_pos_left (a := a) (b := b) (c := c)
33+
rw [PartialOrder.le_iff_lt_or_eq] at h
34+
cases h with
35+
| inl h => exact Preorder.le_of_lt (p h h')
36+
| inr h => subst h; exact Preorder.le_refl (c * a)
37+
| inr h' => subst h'; simp [Semiring.zero_mul, Preorder.le_refl]
38+
39+
theorem mul_le_mul_of_nonneg_right {a b c : R} (h : a ≤ b) (h' : 0 ≤ c) : a * c ≤ b * c := by
40+
rw [PartialOrder.le_iff_lt_or_eq] at h'
41+
cases h' with
42+
| inl h' =>
43+
have p := mul_lt_mul_of_pos_right (a := a) (b := b) (c := c)
44+
rw [PartialOrder.le_iff_lt_or_eq] at h
45+
cases h with
46+
| inl h => exact Preorder.le_of_lt (p h h')
47+
| inr h => subst h; exact Preorder.le_refl (a * c)
48+
| inr h' => subst h'; simp [Semiring.mul_zero, Preorder.le_refl]
49+
50+
theorem mul_nonneg {a b : R} (h₁ : 0 ≤ a) (h₂ : 0 ≤ b) : 0 ≤ a * b := by
51+
simpa [Semiring.zero_mul] using mul_le_mul_of_nonneg_right h₁ h₂
52+
53+
theorem mul_pos {a b : R} (h₁ : 0 < a) (h₂ : 0 < b) : 0 < a * b := by
54+
simpa [Semiring.zero_mul] using mul_lt_mul_of_pos_right h₁ h₂
55+
56+
theorem sq_nonneg {a : R} (h : 0 ≤ a) : 0 ≤ a^2 := by
57+
rw [Semiring.pow_two]
58+
apply mul_nonneg h h
59+
60+
theorem sq_pos {a : R} (h : 0 < a) : 0 < a^2 := by
61+
rw [Semiring.pow_two]
62+
apply mul_pos h h
63+
64+
end Ring.IsOrdered
65+
66+
end Lean.Grind

0 commit comments

Comments
 (0)