Skip to content

Commit 2b7b966

Browse files
committed
chore: examples
1 parent 3ce5097 commit 2b7b966

File tree

1 file changed

+251
-0
lines changed

1 file changed

+251
-0
lines changed

tests/lean/run/grind_guide_2.lean

Lines changed: 251 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,251 @@
1+
2+
/-
3+
New features since last `grind_guide.lean`
4+
-/
5+
set_option warn.sorry false
6+
open Lean Grind Std
7+
8+
/-!
9+
Complete procedure for linear integer arithmetic
10+
-/
11+
example (x y : Int) :
12+
2711*x + 13*y →
13+
11*x + 13*y ≤ 45
14+
-107*x - 9*y →
15+
7*x - 9*y ≤ 4 → False := by
16+
grind
17+
18+
/-!
19+
The linear integer arithmetic module is parametrized by the `ToInt` type classes.
20+
21+
Optimized `Nat` encoding (this quarter).
22+
-/
23+
#check ToInt
24+
25+
example (a b c : UInt64) : a ≤ 2 → b ≤ 3 → c - a - b = 0 → c ≤ 5 := by
26+
grind
27+
28+
example (a b : Fin 15) : a = 0 → b = 1 → a + b > 2 → False := by
29+
grind
30+
31+
/-!
32+
The commutative ring module is parametrized by several type classes
33+
`CommRing`, `Ring`, `CommSemiring`, `Semiring`, `Field`,
34+
`AddRightCancel`, `NoNatZeroDivisors`, `IsCharP`
35+
-/
36+
example [CommRing α] (a b c : α)
37+
: a + b + c = 3
38+
a^2 + b^2 + c^2 = 5
39+
a^3 + b^3 + c^3 = 7
40+
a^4 + b^4 + c^4 = 9 := by
41+
grobner
42+
43+
example (x : BitVec 8) : (x + 16)*(x - 16) = x^2 := by
44+
grind
45+
46+
/-!
47+
The linear arithmetic module supports `IntModule`, `CommRing`, etc.
48+
-/
49+
50+
-- **Note**: It is just a preorder.
51+
example [CommRing α] [LE α] [LT α] [LawfulOrderLT α] [IsPreorder α] [OrderedRing α] (a b c d e : α) :
52+
2*a + b ≥ 1 → b ≥ 0 → c ≥ 0 → d ≥ 0 → e ≥ 0
53+
→ a ≥ 3*c → c ≥ 6*e → d - e*50
54+
→ a + b + 3*c + d + 2*e < 0 → False := by
55+
grind
56+
57+
example [IntModule α] [LE α] [LT α] [LawfulOrderLT α] [IsLinearOrder α] [OrderedAdd α] (f : α → α) (x : α)
58+
: Zero.zero ≤ x → x ≤ 0 → f x = a → f 0 = a := by
59+
grind
60+
61+
/-!
62+
**Several new features implemented this quarter.**
63+
64+
Performance improvements: normalization, canonicalization, proof generation,
65+
proof trimming, etc.
66+
2.5x faster
67+
-/
68+
69+
open Linarith in
70+
set_option trace.grind.debug.proof true in -- Context should contain only `f 2` and `One`
71+
example [CommRing α] [LE α] [LT α] [LawfulOrderLT α] [IsLinearOrder α] [OrderedRing α] (f : Nat → α) :
72+
f 1 <= 0 → f 2 <= 0 → f 3 <= 0 → f 4 <= 0 → f 5 <= 0 → f 6 <= 0 → f 7 <= 0 → f 8 <= 0 → -1 * f 2 + 1 <= 0 → False := by
73+
grind -order
74+
75+
/-!
76+
Implemented support for `NatModule`
77+
-/
78+
section
79+
variable (M : Type) [NatModule M]
80+
81+
example [AddRightCancel M] (x y : M) : 2 • x + 3 • y + x = 3 • (x + y) := by
82+
grind
83+
84+
example [LE M] [LT M] [LawfulOrderLT M] [IsLinearOrder M] [OrderedAdd M] {x y : M}
85+
: x ≤ y → 2 • x + y ≤ 3 • y := by
86+
grind
87+
end
88+
89+
/-!
90+
Implemented normalizers for non-commutative rings and semirings.
91+
-/
92+
example (R : Type u) [Ring R] (a b : R)
93+
: (a - 2 * b)^2 = a^2 - 2 * a * b - 2 * b * a + 4 * b^2 := by
94+
grind
95+
96+
example (R : Type u) [Semiring R] (a b : R)
97+
: (a + 2 * b)^2 = a^2 + 2 * a * b + 2 * b * a + 4 * b^2 := by
98+
grind
99+
100+
/-!
101+
Implemented AC solver.
102+
It is parametrized by the type classes Associative, Commutative, IsLawfulIdentity, IdempotentOp
103+
-/
104+
105+
example {α} (op : α → α → α) [Associative op] (a b c d : α)
106+
: op a b = c →
107+
op b a = d →
108+
op (op c a) (op b c) = op (op a d) (op d b) := by
109+
grind
110+
111+
example {α β : Sort u} (bar : α → β) (op : α → α → α)
112+
[Associative op] [IdempotentOp op] (a b c d e f x y w : α)
113+
: op d (op x c) = op a b →
114+
op e (op f (op y w)) = op (op d a) (op b c) →
115+
bar (op d (op x c)) = bar (op e (op f (op y w))) := by
116+
grind only
117+
118+
/-!
119+
Implemented support for "symbolic" `Fin` and `BitVec`
120+
-/
121+
122+
example (p d : Nat) (n : Fin (p + 1))
123+
: 2 ≤ p → p ≤ d + 1 → d = 1 → n = 0 ∨ n = 1 ∨ n = 2 := by
124+
grind
125+
126+
example {n m : Nat} (x : BitVec n)
127+
: 2 ≤ n → n ≤ m → m = 2 → x = 0 ∨ x = 1 ∨ x = 2 ∨ x = 3 := by
128+
grind
129+
130+
/-!
131+
Implemented bridge between linear and nonlinear solvers.
132+
-/
133+
134+
example (h : s = 4) : 4 < s - 1 + (s - 1) * (s - 1 - 1) / 2 := by
135+
grind
136+
137+
example (a : Nat) (ha : a < 8) (b c : Nat)
138+
: 2 ≤ b → c = 1 → b ≤ c + 1 → a * b < 8 * b := by
139+
grind
140+
141+
/-!
142+
Generalized offset module: `grind order`.
143+
- Support for `Ring` offset constraints (e.g., `x ≤ y + k`)
144+
- Support for any type that implements at least a preorder. (**Extra**)
145+
It is a forward reasoning solver, and computes the transitive closure of this kind of constraint.
146+
It minimizes the number of case splits.
147+
-/
148+
149+
example [LE α] [IsPreorder α]
150+
(a b c : α) : a ≤ b → b ≤ c → a ≤ c := by
151+
grind
152+
153+
example (a b : Int) (h : a + b > 5) : (if a + b ≤ 0 then b else a) = a := by
154+
grind -linarith -cutsat (splits := 0)
155+
156+
example [LE α] [LT α] [LawfulOrderLT α] [IsPreorder α] [Ring α] [OrderedRing α]
157+
(a b : α) : a ≤ 5 → b ≤ 8 → a > 6 ∨ b > 10 → False := by
158+
grind -linarith (splits := 0)
159+
160+
example [LE α] [IsPartialOrder α]
161+
(a b c : α) : a ≤ b → b ≤ c → c ≤ a → a = c := by
162+
grind -linarith
163+
164+
example [LE α] [Std.IsLinearPreorder α]
165+
(a b c d : α) : a ≤ b → ¬ (c ≤ b) → ¬ (d ≤ c) → ¬ (a ≤ d) → False := by
166+
grind -linarith (splits := 0)
167+
168+
/-!
169+
Implemented support injective functions. (**Extra**)
170+
-/
171+
172+
structure InjFn (α : Type) (β : Type) where
173+
f : α → β
174+
h : Function.Injective f
175+
176+
instance : CoeFun (InjFn α β) (fun _ => α → β) where
177+
coe s := s.f
178+
179+
@[grind inj] theorem fn_inj (F : InjFn α β) : Function.Injective (F : α → β) := by
180+
grind [Function.Injective, cases InjFn]
181+
182+
def toList (a : α) : List α := [a]
183+
184+
@[grind inj] theorem toList_inj : Function.Injective (toList : α → List α) := by
185+
grind [Function.Injective, toList]
186+
187+
def succ (x : Nat) := x+1
188+
189+
@[grind inj] theorem succ_inj : Function.Injective succ := by
190+
grind [Function.Injective, succ]
191+
192+
example (F : InjFn α Nat) : toList (succ (F x)) = a → a = toList (succ (F y)) → x = y := by
193+
grind
194+
195+
/-!
196+
Code actions for `grind` attributes and revised modifier semantics. (**Extra**)
197+
Updated and documented all modifiers.
198+
199+
`!` modifier for enabling the "minimal indexable subexpression" condition.
200+
-/
201+
202+
opaque p : Nat → Prop
203+
opaque f : Nat → Nat
204+
opaque r : Nat → Nat → Nat
205+
206+
@[grind] theorem rf_eq : p (f x) → r x (f y) = y := sorry
207+
208+
/-!
209+
Improved diagnostics based on feedback by Bhavik (**Extra**)
210+
-/
211+
212+
example {xs : List α} {i : Nat} (h : i < xs.length) :
213+
xs.take i ++ xs[i] :: xs.drop (i+1) = ys := by
214+
apply List.ext_getElem
215+
next => sorry
216+
next => sorry -- grind
217+
218+
/-!
219+
Implemented solver plugin API (**Extra**)
220+
221+
See: `Order.lean`
222+
-/
223+
224+
/-!
225+
Annotation analyzer (**Extra**)
226+
It will become a new command next quarter.
227+
-/
228+
229+
/-
230+
Next quarter:
231+
- AC E-matching
232+
- ungrind: `grind` tactic mode.
233+
234+
example : ... := by
235+
grind => -- Enters `grind` tactic mode like we have for `conv`.
236+
-- When we hover terms and facts in the Info View we get see the anchors ;:
237+
-- #f124 : ∀ {a b c}, lt b a = false → lt c b = false → lt c a = false
238+
-- #d0a9 : lt xs[j] xs[i] = false → lt xs[i] xs[i] = false
239+
-- #p7ce : lt xs[j] xs[i] = true ∨ j ≤ i
240+
241+
-- Attach a multi-pattern to the transitivity lemma (by anchor)
242+
pattern #f124 => lt b a = false, lt c b = false
243+
244+
-- Instantiate two useful ∀-facts (arguments may be anchors or terms)
245+
instantiate #f124 xs[i] xs[j] xs[i]
246+
247+
-- Case split on a propositional disjunction, name the new hypotheses
248+
cases #p7ce with hLt hLe
249+
next => lia -- close the goal using cutsat
250+
next => finalize -- close the goal using the current search strategy used in `grind`
251+
-/

0 commit comments

Comments
 (0)