Skip to content

Commit 02c5700

Browse files
authored
feat: change apply_rfl tactic so that it does not operate on = (#3784)
Previously: If the `rfl` macro was going to fail, it would: 1. expand to `eq_refl`, which is implemented by `Lean.Elab.Tactic.evalRefl`, and call `Lean.MVarId.refl` which would: * either try kernel defeq (if in `.default` or `.all` transparency mode) * otherwise try `IsDefEq` * then fail. 2. Next expand to the `apply_rfl` tactic, which is implemented by `Lean.Elab.Tactic.Rfl.evalApplyRfl`, and call `Lean.MVarId.applyRefl` which would look for lemmas labelled `@[refl]`, and unfortunately in Mathlib find `Eq.refl`, so try applying that (resulting in another `IsDefEq`) 3. Because of an accidental duplication, if `Lean.Elab.Tactic.Rfl` was imported, it would *again* expand to `apply_rfl`. Now: 1. Same behaviour in `eq_refl`. 2. The `@[refl]` attribute will reject `Eq.refl`, and `MVarId.applyRefl` will fail when applied to equality goals. 3. The duplication has been removed.
1 parent 3ee1cdf commit 02c5700

File tree

4 files changed

+29
-30
lines changed

4 files changed

+29
-30
lines changed

src/Init/Tactics.lean

Lines changed: 3 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -377,8 +377,9 @@ macro_rules | `(tactic| rfl) => `(tactic| eq_refl)
377377
macro_rules | `(tactic| rfl) => `(tactic| exact HEq.rfl)
378378

379379
/--
380-
This tactic applies to a goal whose target has the form `x ~ x`, where `~` is a reflexive
381-
relation, that is, a relation which has a reflexive lemma tagged with the attribute [refl].
380+
This tactic applies to a goal whose target has the form `x ~ x`,
381+
where `~` is a reflexive relation other than `=`,
382+
that is, a relation which has a reflexive lemma tagged with the attribute @[refl].
382383
-/
383384
syntax (name := applyRfl) "apply_rfl" : tactic
384385

src/Lean/Elab/Tactic/Rfl.lean

Lines changed: 1 addition & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -10,19 +10,12 @@ import Lean.Elab.Tactic.Basic
1010
/-!
1111
# `rfl` tactic extension for reflexive relations
1212
13-
This extends the `rfl` tactic so that it works on any reflexive relation,
13+
This extends the `rfl` tactic so that it works on reflexive relations other than `=`,
1414
provided the reflexivity lemma has been marked as `@[refl]`.
1515
-/
1616

1717
namespace Lean.Elab.Tactic.Rfl
1818

19-
/--
20-
This tactic applies to a goal whose target has the form `x ~ x`, where `~` is a reflexive
21-
relation, that is, a relation which has a reflexive lemma tagged with the attribute [refl].
22-
-/
23-
elab_rules : tactic
24-
| `(tactic| rfl) => withMainContext do liftMetaFinishingTactic (·.applyRfl)
25-
2619
@[builtin_tactic Lean.Parser.Tactic.applyRfl] def evalApplyRfl : Tactic := fun stx =>
2720
match stx with
2821
| `(tactic| apply_rfl) => withMainContext do liftMetaFinishingTactic (·.applyRfl)

src/Lean/Meta/Tactic/Rfl.lean

Lines changed: 25 additions & 18 deletions
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,7 @@ Authors: Newell Jensen, Thomas Murrills
66
prelude
77
import Lean.Meta.Tactic.Apply
88
import Lean.Elab.Tactic.Basic
9+
import Lean.Meta.Tactic.Refl
910

1011
/-!
1112
# `rfl` tactic extension for reflexive relations
@@ -38,6 +39,8 @@ initialize registerBuiltinAttribute {
3839
let fail := throwError
3940
"@[refl] attribute only applies to lemmas proving x ∼ x, got {declTy}"
4041
let .app (.app rel lhs) rhs := targetTy | fail
42+
if let .app (.const ``Eq [_]) _ := rel then
43+
throwError "@[refl] attribute may not be used on `Eq.refl`."
4144
unless ← withNewMCtxDepth <| isDefEq lhs rhs do fail
4245
let key ← DiscrTree.mkPath rel reflExt.config
4346
reflExt.add (decl, key) kind
@@ -47,29 +50,33 @@ open Elab Tactic
4750

4851
/-- `MetaM` version of the `rfl` tactic.
4952
50-
This tactic applies to a goal whose target has the form `x ~ x`, where `~` is a reflexive
51-
relation, that is, a relation which has a reflexive lemma tagged with the attribute [refl].
53+
This tactic applies to a goal whose target has the form `x ~ x`,
54+
where `~` is a reflexive relation other than `=`,
55+
that is, a relation which has a reflexive lemma tagged with the attribute @[refl].
5256
-/
5357
def _root_.Lean.MVarId.applyRfl (goal : MVarId) : MetaM Unit := do
5458
let .app (.app rel _) _ ← whnfR <|← instantiateMVars <|← goal.getType
5559
| throwError "reflexivity lemmas only apply to binary relations, not{
5660
indentExpr (← goal.getType)}"
57-
let s ← saveState
58-
let mut ex? := none
59-
for lem in ← (reflExt.getState (← getEnv)).getMatch rel reflExt.config do
60-
try
61-
let gs ← goal.apply (← mkConstWithFreshMVarLevels lem)
62-
if gs.isEmpty then return () else
63-
logError <| MessageData.tagged `Tactic.unsolvedGoals <| m!"unsolved goals\n{
64-
goalsToMessageData gs}"
65-
catch e =>
66-
ex? := ex? <|> (some (← saveState, e)) -- stash the first failure of `apply`
67-
s.restore
68-
if let some (sErr, e) := ex? then
69-
sErr.restore
70-
throw e
61+
if let .app (.const ``Eq [_]) _ := rel then
62+
throwError "MVarId.applyRfl does not solve `=` goals. Use `MVarId.refl` instead."
7163
else
72-
throwError "rfl failed, no lemma with @[refl] applies"
64+
let s ← saveState
65+
let mut ex? := none
66+
for lem in ← (reflExt.getState (← getEnv)).getMatch rel reflExt.config do
67+
try
68+
let gs ← goal.apply (← mkConstWithFreshMVarLevels lem)
69+
if gs.isEmpty then return () else
70+
logError <| MessageData.tagged `Tactic.unsolvedGoals <| m!"unsolved goals\n{
71+
goalsToMessageData gs}"
72+
catch e =>
73+
ex? := ex? <|> (some (← saveState, e)) -- stash the first failure of `apply`
74+
s.restore
75+
if let some (sErr, e) := ex? then
76+
sErr.restore
77+
throw e
78+
else
79+
throwError "rfl failed, no lemma with @[refl] applies"
7380

7481
/-- Helper theorem for `Lean.MVarId.liftReflToEq`. -/
7582
private theorem rel_of_eq_and_refl {α : Sort _} {R : α → α → Prop}
@@ -78,7 +85,7 @@ private theorem rel_of_eq_and_refl {α : Sort _} {R : α → α → Prop}
7885

7986
/--
8087
Convert a goal of the form `x ~ y` into the form `x = y`, where `~` is a reflexive
81-
relation, that is, a relation which has a reflexive lemma tagged with the attribute `[refl]`.
88+
relation, that is, a relation which has a reflexive lemma tagged with the attribute `@[refl]`.
8289
If this can't be done, returns the original `MVarId`.
8390
-/
8491
def _root_.Lean.MVarId.liftReflToEq (mvarId : MVarId) : MetaM MVarId := do

tests/lean/run/rewrites.lean

Lines changed: 0 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1,5 +1,3 @@
1-
attribute [refl] Eq.refl
2-
31
private axiom test_sorry : ∀ {α}, α
42

53
-- To see the (sorted) list of lemmas that `rw?` will try rewriting by, use:

0 commit comments

Comments
 (0)