Skip to content

Commit 42c1a72

Browse files
committed
feat: rename get_elem_tactic_trivial to get_elem_tactic_extensible
1 parent 2a1354b commit 42c1a72

File tree

5 files changed

+28
-12
lines changed

5 files changed

+28
-12
lines changed

src/Init/Data/Range/Basic.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -84,5 +84,8 @@ theorem Membership.mem.step {i : Nat} {r : Std.Range} (h : i ∈ r) : (i - r.sta
8484
theorem Membership.get_elem_helper {i n : Nat} {r : Std.Range} (h₁ : i ∈ r) (h₂ : r.stop = n) :
8585
i < n := h₂ ▸ h₁.2.1
8686

87+
macro_rules
88+
| `(tactic| get_elem_tactic_extensible) => `(tactic| apply Membership.get_elem_helper; assumption; rfl)
89+
8790
macro_rules
8891
| `(tactic| get_elem_tactic_trivial) => `(tactic| apply Membership.get_elem_helper; assumption; rfl)

src/Init/GetElem.lean

Lines changed: 4 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -47,7 +47,7 @@ proof in the context using `have`, because `get_elem_tactic` tries
4747
4848
The proof side-condition `valid xs i` is automatically dispatched by the
4949
`get_elem_tactic` tactic; this tactic can be extended by adding more clauses to
50-
`get_elem_tactic_trivial` using `macro_rules`.
50+
`get_elem_tactic_extensible` using `macro_rules`.
5151
5252
`xs[i]?` and `xs[i]!` do not impose a proof obligation; the former returns
5353
an `Option elem`, with `none` signalling that the value isn't present, and
@@ -282,6 +282,9 @@ instance [GetElem? cont Nat elem dom] [h : LawfulGetElem cont Nat elem dom] :
282282

283283
@[simp] theorem getElem!_fin [GetElem? Cont Nat Elem Dom] (a : Cont) (i : Fin n) [Inhabited Elem] : a[i]! = a[i.1]! := rfl
284284

285+
macro_rules
286+
| `(tactic| get_elem_tactic_extensible) => `(tactic| (with_reducible apply Fin.val_lt_of_le); get_elem_tactic_extensible; done)
287+
285288
macro_rules
286289
| `(tactic| get_elem_tactic_trivial) => `(tactic| (with_reducible apply Fin.val_lt_of_le); get_elem_tactic_trivial; done)
287290

src/Init/Tactics.lean

Lines changed: 17 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -1880,30 +1880,37 @@ syntax "‹" withoutPosition(term) "›" : term
18801880
macro_rules | `(‹$type›) => `((by assumption : $type))
18811881

18821882
/--
1883-
`get_elem_tactic_trivial` is an extensible tactic automatically called
1883+
`get_elem_tactic_extensible` is an extensible tactic automatically called
18841884
by the notation `arr[i]` to prove any side conditions that arise when
18851885
constructing the term (e.g. the index is in bounds of the array).
1886-
The default behavior is to just try `trivial` (which handles the case
1887-
where `i < arr.size` is in the context) and `simp +arith` and `omega`
1886+
The default behavior is to try `simp +arith` and `omega`
18881887
(for doing linear arithmetic in the index).
1888+
1889+
(Note that the core tactic `get_elem_tactic` has already tried
1890+
`done` and `assumption` before the extensible tactic is called.)
18891891
-/
1892+
syntax "get_elem_tactic_extensible" : tactic
1893+
1894+
/-- `get_elem_tactic_trivial` has been deprecated in favour of `get_elem_tactic_extensible`. -/
18901895
syntax "get_elem_tactic_trivial" : tactic
18911896

1897+
macro_rules | `(tactic| get_elem_tactic_extensible) => `(tactic| omega)
1898+
macro_rules | `(tactic| get_elem_tactic_extensible) => `(tactic| simp +arith; done)
1899+
18921900
macro_rules | `(tactic| get_elem_tactic_trivial) => `(tactic| omega)
18931901
macro_rules | `(tactic| get_elem_tactic_trivial) => `(tactic| simp +arith; done)
1894-
macro_rules | `(tactic| get_elem_tactic_trivial) => `(tactic| trivial)
18951902

18961903
/--
18971904
`get_elem_tactic` is the tactic automatically called by the notation `arr[i]`
18981905
to prove any side conditions that arise when constructing the term
18991906
(e.g. the index is in bounds of the array). It just delegates to
1900-
`get_elem_tactic_trivial` and gives a diagnostic error message otherwise;
1901-
users are encouraged to extend `get_elem_tactic_trivial` instead of this tactic.
1907+
`get_elem_tactic_extensible` and gives a diagnostic error message otherwise;
1908+
users are encouraged to extend `get_elem_tactic_extensible` instead of this tactic.
19021909
-/
19031910
macro "get_elem_tactic" : tactic =>
19041911
`(tactic| first
19051912
/-
1906-
Recall that `macro_rules` (namely, for `get_elem_tactic_trivial`) are tried in reverse order.
1913+
Recall that `macro_rules` (namely, for `get_elem_tactic_extensible`) are tried in reverse order.
19071914
We first, however, try `done`, since the necessary proof may already have been
19081915
found during unification, in which case there is no goal to solve (see #6999).
19091916
If a goal is present, we want `assumption` to be tried first.
@@ -1916,15 +1923,15 @@ macro "get_elem_tactic" : tactic =>
19161923
If `omega` is used to "fill" this proof, we will have a more complex proof term that
19171924
cannot be inferred by unification.
19181925
We hardcoded `assumption` here to ensure users cannot accidentally break this IF
1919-
they add new `macro_rules` for `get_elem_tactic_trivial`.
1926+
they add new `macro_rules` for `get_elem_tactic_extensible`.
19201927
19211928
TODO: Implement priorities for `macro_rules`.
1922-
TODO: Ensure we have **high-priority** macro_rules for `get_elem_tactic_trivial` which are
1929+
TODO: Ensure we have **high-priority** macro_rules for `get_elem_tactic_extensible` which are
19231930
just `done` and `assumption`.
19241931
-/
19251932
| done
19261933
| assumption
1927-
| get_elem_tactic_trivial
1934+
| get_elem_tactic_extensible
19281935
| fail "failed to prove index is valid, possible solutions:
19291936
- Use `have`-expressions to prove the index is valid
19301937
- Use `a[i]!` notation instead, runtime check is performed, and 'Panic' error message is produced if index is not valid

src/Lean/Elab/Tactic.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -51,3 +51,6 @@ import Lean.Elab.Tactic.TreeTacAttr
5151
import Lean.Elab.Tactic.ExposeNames
5252
import Lean.Elab.Tactic.SimpArith
5353
import Lean.Elab.Tactic.Lets
54+
55+
elab "get_elem_tactic_trivial" : tactic => do
56+
throwError "`get_elem_tactic_trivial` has been renamed to `get_elem_tactic_extensible`"

tests/lean/run/grind_heapsort.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ set_option grind.warning false
44
/-
55
Use `grind` as one of the tactics for array-element access and termination proofs.
66
-/
7-
macro_rules | `(tactic| get_elem_tactic_trivial) => `(tactic| grind)
7+
macro_rules | `(tactic| get_elem_tactic_extensible) => `(tactic| grind)
88
/-
99
Note: We disable model-based theory combination (-mbtc) here because `grind` can
1010
exhaust heartbeats when exploring certain "bad" termination checker scenarios.

0 commit comments

Comments
 (0)