Skip to content

Commit c447793

Browse files
authored
feat: Int.subNatNat in grind (#11615)
This PR adds a normalization rule for `Int.subNatNat` to `grind`. Closes #11543
1 parent 864acdd commit c447793

File tree

2 files changed

+179
-1
lines changed

2 files changed

+179
-1
lines changed

src/Init/Grind/Norm.lean

Lines changed: 4 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -149,6 +149,9 @@ theorem smul_nat_eq_mul {α} [Semiring α] (n : Nat) (a : α) : n • a = NatCas
149149
theorem smul_int_eq_mul {α} [Ring α] (i : Int) (a : α) : i • a = Int.cast i * a := by
150150
rw [Ring.zsmul_eq_intCast_mul]
151151

152+
theorem Int.subNatNat_eq (a b : Nat) : Int.subNatNat a b = NatCast.natCast a - NatCast.natCast b := by
153+
apply Int.subNatNat_eq_coe
154+
152155
-- Remark: for additional `grind` simprocs, check `Lean/Meta/Tactic/Grind`
153156
init_grind_norm
154157
/- Pre theorems -/
@@ -188,7 +191,7 @@ init_grind_norm
188191
natCast_div natCast_mod natCast_id
189192
natCast_add natCast_mul natCast_pow
190193
Int.one_pow
191-
Int.pow_zero Int.pow_one
194+
Int.pow_zero Int.pow_one Int.subNatNat_eq
192195
-- Int op folding
193196
Int.add_def Int.mul_def Int.ofNat_eq_coe
194197
Int.Linear.sub_fold Int.Linear.neg_fold

tests/lean/run/gring_11543.lean

Lines changed: 175 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,175 @@
1+
section Mathlib.Algebra.Group.Defs
2+
3+
class Monoid (M : Type u) extends Mul M where
4+
5+
class AddGroup (G : Type u) extends Add G, Neg G, Sub G where
6+
7+
theorem add_sub_assoc {G : Type} [AddGroup G] (a b c : G) : a + b - c = a + (b - c) := by
8+
sorry
9+
10+
end Mathlib.Algebra.Group.Defs
11+
12+
abbrev Function.swap {φ : α → β → Sort u₃} (f : ∀ x y, φ x y) : ∀ y x, φ x y := fun y x => f x y
13+
14+
section Mathlib.Algebra.Group.Basic
15+
16+
@[simp]
17+
theorem add_sub_cancel {G} [AddGroup G](a b : G) : a + (b - a) = b := by
18+
sorry
19+
20+
end Mathlib.Algebra.Group.Basic
21+
22+
section Mathlib.Algebra.Order.Monoid.Unbundled.Defs
23+
24+
open Function
25+
26+
variable (M N : Type) (μ : M → N → N) (r : N → N → Prop)
27+
28+
def Covariant : Prop :=
29+
∀ (m) {n₁ n₂}, r n₁ n₂ → r (μ m n₁) (μ m n₂)
30+
31+
class CovariantClass : Prop where
32+
protected elim : Covariant M N μ r
33+
34+
abbrev AddLeftMono [Add M] [LE M] : Prop :=
35+
CovariantClass M M (· + ·) (· ≤ ·)
36+
37+
abbrev AddRightMono [Add M] [LE M] : Prop :=
38+
CovariantClass M M (swap (· + ·)) (· ≤ ·)
39+
40+
instance covariant_swap_add_of_covariant_add [Add N]
41+
[CovariantClass N N (· + ·) r] : CovariantClass N N (swap (· + ·)) r where
42+
elim := sorry
43+
44+
end Mathlib.Algebra.Order.Monoid.Unbundled.Defs
45+
46+
section Mathlib.Algebra.Order.Monoid.Unbundled.Basic
47+
48+
instance Int.instAddLeftMono : AddLeftMono Int where
49+
elim := sorry
50+
51+
end Mathlib.Algebra.Order.Monoid.Unbundled.Basic
52+
53+
section Mathlib.Algebra.Order.Sub.Defs
54+
55+
class OrderedSub (α : Type) [LE α] [Add α] [Sub α] : Prop where
56+
57+
@[simp]
58+
theorem tsub_le_iff_right {α : Type} [LE α] [Add α] [Sub α] [OrderedSub α] {a b c : α} :
59+
a - b ≤ c ↔ a ≤ c + b := sorry
60+
61+
end Mathlib.Algebra.Order.Sub.Defs
62+
63+
section Mathlib.Algebra.Order.Monoid.Unbundled.Basic
64+
65+
instance AddGroup.toOrderedSub {α : Type} [AddGroup α] [LE α]
66+
[AddRightMono α] : OrderedSub α := sorry
67+
68+
end Mathlib.Algebra.Order.Monoid.Unbundled.Basic
69+
70+
section Mathlib.Data.Nat.Cast.Defs
71+
72+
class AddMonoidWithOne (R : Type) extends NatCast R, Add R where
73+
74+
theorem Nat.cast_add [AddMonoidWithOne R] (m n : Nat) : ((m + n : Nat) : R) = m + n := sorry
75+
76+
end Mathlib.Data.Nat.Cast.Defs
77+
78+
section Mathlib.Data.Int.Cast.Defs
79+
80+
class AddGroupWithOne (R : Type) extends IntCast R, AddMonoidWithOne R, AddGroup R where
81+
82+
end Mathlib.Data.Int.Cast.Defs
83+
84+
section Mathlib.Data.Int.Cast.Basic
85+
86+
namespace Int
87+
88+
variable {R : Type} [AddGroupWithOne R]
89+
90+
@[norm_cast]
91+
theorem cast_subNatNat (m n) : ((Int.subNatNat m n : Int) : R) = m - n := sorry
92+
93+
end Int
94+
95+
end Mathlib.Data.Int.Cast.Basic
96+
97+
section Mathlib.Algebra.Group.Int.Defs
98+
99+
namespace Int
100+
101+
instance instCommMonoid : Monoid Int where
102+
103+
instance instAddCommGroup : AddGroup Int where
104+
105+
instance instCommSemigroup : Mul Int := by infer_instance
106+
107+
end Int
108+
109+
end Mathlib.Algebra.Group.Int.Defs
110+
111+
section Mathlib.Algebra.Ring.Defs
112+
113+
class NonAssocSemiring (α : Type) extends Add α, Mul α, Zero α
114+
115+
class NonAssocRing (α : Type) extends AddGroup α, NonAssocSemiring α
116+
117+
class Ring (R : Type) extends NonAssocSemiring R, AddGroupWithOne R
118+
119+
section NonAssocRing
120+
121+
variable {α : Type} [NonAssocRing α]
122+
123+
theorem mul_sub_left_distrib (a b c : α) : a * (b - c) = a * b - a * c := sorry
124+
125+
end NonAssocRing
126+
127+
section Ring
128+
129+
variable {α : Type} [Ring α]
130+
131+
instance (priority := 100) Ring.toNonAssocRing : NonAssocRing α :=
132+
{ ‹Ring α› with }
133+
134+
end Ring
135+
136+
class CommRing (α : Type) extends Ring α, Monoid α
137+
138+
end Mathlib.Algebra.Ring.Defs
139+
140+
section Mathlib.Algebra.Ring.Int.Defs
141+
142+
namespace Int
143+
144+
instance instCommRing : CommRing Int where
145+
intCast := (·)
146+
147+
end Int
148+
149+
end Mathlib.Algebra.Ring.Int.Defs
150+
151+
example {a b : Nat}
152+
(this : (b : Int) + ↑b * (↑a - ↑b) ≤ (↑b + (↑a - ↑b))) :
153+
a * b ≤ a + b * b := by
154+
simp [mul_sub_left_distrib, ← add_sub_assoc] at this
155+
norm_cast at this
156+
grind
157+
158+
/--
159+
info: Try this:
160+
[apply] simp only [mul_sub_left_distrib, ← add_sub_assoc, add_sub_cancel, tsub_le_iff_right] at this
161+
-/
162+
#guard_msgs in
163+
example {a b : Nat}
164+
(this : (b : Int) + ↑b * (↑a - ↑b) ≤ (↑b + (↑a - ↑b))) :
165+
a * b ≤ a + b * b := by
166+
simp? [mul_sub_left_distrib, ← add_sub_assoc] at this
167+
norm_cast at this
168+
grind
169+
170+
example {a b : Nat}
171+
(this : (b : Int) + ↑b * (↑a - ↑b) ≤ (↑b + (↑a - ↑b))) :
172+
a * b ≤ a + b * b := by
173+
simp only [mul_sub_left_distrib, ← add_sub_assoc, add_sub_cancel, tsub_le_iff_right] at this
174+
norm_cast at this
175+
grind

0 commit comments

Comments
 (0)