-
Notifications
You must be signed in to change notification settings - Fork 546
/
Copy pathRfl.lean
110 lines (97 loc) · 4.05 KB
/
Rfl.lean
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
/-
Copyright (c) 2022 Newell Jensen. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Newell Jensen, Thomas Murrills
-/
prelude
import Lean.Meta.Tactic.Apply
import Lean.Elab.Tactic.Basic
import Lean.Meta.Tactic.Refl
/-!
# `rfl` tactic extension for reflexive relations
This extends the `rfl` tactic so that it works on any reflexive relation,
provided the reflexivity lemma has been marked as `@[refl]`.
-/
namespace Lean.Meta.Rfl
open Lean Meta
/-- Discrimation tree settings for the `refl` extension. -/
def reflExt.config : WhnfCoreConfig := {}
/-- Environment extensions for `refl` lemmas -/
initialize reflExt :
SimpleScopedEnvExtension (Name × Array DiscrTree.Key) (DiscrTree Name) ←
registerSimpleScopedEnvExtension {
addEntry := fun dt (n, ks) => dt.insertCore ks n
initial := {}
}
initialize registerBuiltinAttribute {
name := `refl
descr := "reflexivity relation"
add := fun decl _ kind => MetaM.run' do
let declTy := (← getConstInfo decl).type
let (_, _, targetTy) ← withReducible <| forallMetaTelescopeReducing declTy
let fail := throwError
"@[refl] attribute only applies to lemmas proving x ∼ x, got {declTy}"
let .app (.app rel lhs) rhs := targetTy | fail
if let .app (.const ``Eq [_]) _ := rel then
throwError "@[refl] attribute may not be used on `Eq.refl`."
unless ← withNewMCtxDepth <| isDefEq lhs rhs do fail
let key ← DiscrTree.mkPath rel reflExt.config
reflExt.add (decl, key) kind
}
open Elab Tactic
/-- `MetaM` version of the `rfl` tactic.
This tactic applies to a goal whose target has the form `x ~ x`,
where `~` is a reflexive relation other than `=`,
that is, a relation which has a reflexive lemma tagged with the attribute @[refl].
-/
def _root_.Lean.MVarId.applyRfl (goal : MVarId) : MetaM Unit := do
let .app (.app rel _) _ ← whnfR <|← instantiateMVars <|← goal.getType
| throwError "reflexivity lemmas only apply to binary relations, not{
indentExpr (← goal.getType)}"
if let .app (.const ``Eq [_]) _ := rel then
throwError "MVarId.applyRfl does not solve `=` goals. Use `MVarId.refl` instead."
else
let s ← saveState
let mut ex? := none
for lem in ← (reflExt.getState (← getEnv)).getMatch rel reflExt.config do
try
let gs ← goal.apply (← mkConstWithFreshMVarLevels lem)
if gs.isEmpty then return () else
logError <| MessageData.tagged `Tactic.unsolvedGoals <| m!"unsolved goals\n{
goalsToMessageData gs}"
catch e =>
ex? := ex? <|> (some (← saveState, e)) -- stash the first failure of `apply`
s.restore
if let some (sErr, e) := ex? then
sErr.restore
throw e
else
throwError "rfl failed, no lemma with @[refl] applies"
/-- Helper theorem for `Lean.MVarId.liftReflToEq`. -/
private theorem rel_of_eq_and_refl {α : Sort _} {R : α → α → Prop}
{x y : α} (hxy : x = y) (h : R x x) : R x y :=
hxy ▸ h
/--
Convert a goal of the form `x ~ y` into the form `x = y`, where `~` is a reflexive
relation, that is, a relation which has a reflexive lemma tagged with the attribute `@[refl]`.
If this can't be done, returns the original `MVarId`.
-/
def _root_.Lean.MVarId.liftReflToEq (mvarId : MVarId) : MetaM MVarId := do
mvarId.checkNotAssigned `liftReflToEq
let .app (.app rel _) _ ← withReducible mvarId.getType' | return mvarId
if rel.isAppOf `Eq then
-- No need to lift Eq to Eq
return mvarId
for lem in ← (reflExt.getState (← getEnv)).getMatch rel reflExt.config do
let res ← observing? do
-- First create an equality relating the LHS and RHS
-- and reduce the goal to proving that LHS is related to LHS.
let [mvarIdEq, mvarIdR] ← mvarId.apply (← mkConstWithFreshMVarLevels ``rel_of_eq_and_refl)
| failure
-- Then fill in the proof of the latter by reflexivity.
let [] ← mvarIdR.apply (← mkConstWithFreshMVarLevels lem) | failure
return mvarIdEq
if let some mvarIdEq := res then
return mvarIdEq
return mvarId
end Lean.Meta.Rfl