Dependent rewrite tactic #
See Config.castMode
.
- proofs : CastMode
Only insert casts on proofs.
In this mode, it is not permitted to cast subterms of proofs that are not themselves proofs.
- all : CastMode
Insert casts whenever necessary.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Equations
- Mathlib.Tactic.DepRewrite.instLECastMode = { le := fun (a b : Mathlib.Tactic.DepRewrite.CastMode) => a.toNat ≤ b.toNat }
Configures the behavior of the rewrite!
and rw!
tactics.
- transparency : Lean.Meta.TransparencyMode
Which transparency level to use when unifying the rewrite rule's LHS against subterms of the term being rewritten.
- occs : Lean.Meta.Occurrences
Which occurrences to rewrite.
- castMode : CastMode
The cast mode specifies when
rw!
is permitted to insert casts in order to correct subterms that become type-incorect as a result of rewriting.For example, given
P : Nat → Prop
,f : (n : Nat) → P n → Nat
andh : P n₀
, rewritingf n₀ h
byeq : n₀ = n₁
producesf n₁ h
, whereh
does not typecheck atP n₁
. The tactic will casth
toeq ▸ h : P n₁
iff.proofs ≤ castMode
.
Instances For
- cfg : Config
Configuration.
- p : Lean.Expr
The pattern to generalize over.
- x : Lean.Expr
The free variable to substitute for
p
. - h : Lean.Expr
- Δ : Array (Lean.FVarId × Lean.Expr)
The list of value-less binders (
cdecl
s and nondependentldecl
s) that we have introduced. Together with each binder, we store its type abstracted overx
andh
, and with all occurrences of previous entries inΔ
casted along the abstracting equation.E.g., if the local context is
a : T, b : U
, we store(a, Ma)
whereMa := fun (x' : α) (h' : x = x') => T[x'/x, h'/h]
and(b, fun (x' : α) (h' : x = x') => U[x'/x, h'/h, (Eq.rec (motive := Ma) a h)/a])
See the docstring onvisitAndCast
. The set of all dependent
ldecl
s that we have introduced.- pHeadIdx : Lean.HeadIndex
Cached
p.toHeadIndex
. - pNumArgs : Nat
Cached
p.toNumArgs
.
Instances For
We use a cache entry iff the upcoming traversal would abstract exactly the same occurrences as the cached traversal.
Equations
- One or more equations did not get rendered due to their size.
- Mathlib.Tactic.DepRewrite.canUseCache cacheOcc dCacheOcc currOcc Lean.Meta.Occurrences.all = true
Instances For
Monad for computing dabstract
.
The Nat
state tracks which occurrence of the pattern we are about to see, 1-indexed
(so the initial value is 1
).
The cache stores results of visit
together with
- the
Nat
state before the cached call; and - the difference in the state resulting from the call.
We store these because even if the cache hits,
we must update the state as if the call had been made.
Storing the difference suffices because the state increases monotonically.
See also
canUseCache
.
Equations
Instances For
Check that casting e : t
is allowed in the current mode.
(We don't need to know what type e
is cast to:
we only check the sort of t
, and it cannot change.)
Equations
- One or more equations did not get rendered due to their size.
Instances For
In e
, inline the values of those ldecl
s that appear in fvars
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
If e : te
is a term whose type mentions x
, h
(the generalization variables)
or entries in Δ
/δ
,
return h.symm ▸ e : te[p/x, rfl/h, …]
.
Otherwise return none
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Compute the motive that casts e
back to te[p/x, rfl/h, …]
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Cast e : te[p/x, rfl/h, ...]
to h ▸ e : te
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Given e
, return e'
where e'
has had
- the occurrences of
p
inctx.cfg.occs
replaced byx
; and - subterms cast as appropriate in order to make
e'
type-correct.
If et?
is not none
, the output is guaranteed to have type (defeq to) et?
.
We do not assume that e
is well-typed.
We use this when processing binders:
to traverse ∀ (x : α), β
,
we obtain α' ← visit α
,
add x : α'
to the local context
and continue traversing β
.
Although x : α' ⊢ β
may not hold,
the output β'
should have x : α' ⊢ β'
(otherwise we have a bug).
To achieve this, we maintain the invariant
that all entries in the local context that we have introduced
can be translated back to their original (pre-visit
) types
using the motive computed by castBack?.motive
.
(But we have not attempted to prove this.)
Like visitAndCast
, but does not insert casts at the top level.
The expected types of certain subterms are computed from et?
.
Analogue of kabstract
with support for inserting casts.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Analogue of Lean.MVarId.rewrite
with support for inserting casts.
Equations
- One or more equations did not get rendered due to their size.
Instances For
rewrite!
is like rewrite
,
but can also insert casts to adjust types that depend on the LHS of a rewrite.
It is available as an ordinary tactic and a conv
tactic.
The sort of casts that are inserted is controlled by the castMode
configuration option.
By default, only proof terms are casted;
by proof irrelevance, this adds no observable complexity.
With rewrite! +letAbs (castMode := .all)
, casts are inserted whenever necessary.
This means that the 'motive is not type correct' error never occurs,
at the expense of creating potentially complicated terms.
Equations
- One or more equations did not get rendered due to their size.
Instances For
rw!
is like rewrite!
, but also calls dsimp
to simplify the result after every substitution.
It is available as an ordinary tactic and a conv
tactic.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Apply rewrite!
to the goal.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Apply rewrite!
to a local declaration.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Elaborate DepRewrite.Config
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
rewrite!
is like rewrite
,
but can also insert casts to adjust types that depend on the LHS of a rewrite.
It is available as an ordinary tactic and a conv
tactic.
The sort of casts that are inserted is controlled by the castMode
configuration option.
By default, only proof terms are casted;
by proof irrelevance, this adds no observable complexity.
With rewrite! +letAbs (castMode := .all)
, casts are inserted whenever necessary.
This means that the 'motive is not type correct' error never occurs,
at the expense of creating potentially complicated terms.
Equations
- One or more equations did not get rendered due to their size.
Instances For
rw!
is like rewrite!
, but also calls dsimp
to simplify the result after every substitution.
It is available as an ordinary tactic and a conv
tactic.
Equations
- One or more equations did not get rendered due to their size.
Instances For
rewrite!
is like rewrite
,
but can also insert casts to adjust types that depend on the LHS of a rewrite.
It is available as an ordinary tactic and a conv
tactic.
The sort of casts that are inserted is controlled by the castMode
configuration option.
By default, only proof terms are casted;
by proof irrelevance, this adds no observable complexity.
With rewrite! +letAbs (castMode := .all)
, casts are inserted whenever necessary.
This means that the 'motive is not type correct' error never occurs,
at the expense of creating potentially complicated terms.
Equations
- One or more equations did not get rendered due to their size.
Instances For
rw!
is like rewrite!
, but also calls dsimp
to simplify the result after every substitution.
It is available as an ordinary tactic and a conv
tactic.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Apply rewrite!
to the goal.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Apply rw!
to the goal.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Apply rw!
to a local declaration.
Equations
- One or more equations did not get rendered due to their size.
Instances For
rewrite!
is like rewrite
,
but can also insert casts to adjust types that depend on the LHS of a rewrite.
It is available as an ordinary tactic and a conv
tactic.
The sort of casts that are inserted is controlled by the castMode
configuration option.
By default, only proof terms are casted;
by proof irrelevance, this adds no observable complexity.
With rewrite! +letAbs (castMode := .all)
, casts are inserted whenever necessary.
This means that the 'motive is not type correct' error never occurs,
at the expense of creating potentially complicated terms.
Equations
- One or more equations did not get rendered due to their size.
Instances For
rw!
is like rewrite!
, but also calls dsimp
to simplify the result after every substitution.
It is available as an ordinary tactic and a conv
tactic.
Equations
- One or more equations did not get rendered due to their size.