congrm
: congr
with pattern-matching
congrm e
gives to the use the functionality of using congr
with an expression e
"guiding"
congr
through the matching. This allows more flexibility than congr' n
, which enters uniformly
through n
iterations. Instead, we can guide the matching deeper on some parts of the expression
and stop earlier on other parts.
Implementation notes #
Function underscores #
See the doc-string to tactic.interactive.congrm
for more details. Here we describe how to add
more "function underscores".
The pattern for generating a function underscore is to define a "generic" n
-ary function, for some
number n
. You can take a look at tactic.congrm_fun_1, ..., tactic.congrm_fun_4
.
These implement the "function underscores" _₁, ..., _₄
. If you want a different arity for your
function, simply
introduce
@[nolint unused_arguments]
def congrm_fun_n {α₁ … αₙ ρ} {r : ρ} : α₁ → ⋯ → aₙ → ρ := λ _ … _, r
notation `_ₙ` := congrm_fun_n
Warning: convert_to_explicit
checks that the first 18 characters in the name of _ₙ
are
identical to tactic.congrm_fun_
to perform its job. Thus, if you want to implement
"function underscores" with different arity, either make sure that their names begin with
tactic.congrm_fun_
or you should change convert_to_explicit
accordingly.
Assume that the goal is of the form lhs = rhs
or lhs ↔ rhs
.
congrm e
takes an expression e
containing placeholders _
and scans e, lhs, rhs
in parallel.
It matches both lhs
and rhs
to the pattern e
, and produces one goal for each placeholder,
stating that the corresponding subexpressions in lhs
and rhs
are equal.
Examples:
example {a b c d : ℕ} :
nat.pred a.succ * (d + (c + a.pred)) = nat.pred b.succ * (b + (c + d.pred)) :=
begin
congrm nat.pred (nat.succ _) * (_ + _),
/- Goals left:
⊢ a = b
⊢ d = b
⊢ c + a.pred = c + d.pred
-/
sorry,
sorry,
sorry,
end
example {a b : ℕ} (h : a = b) : (λ y : ℕ, ∀ z, a + a = z) = (λ x, ∀ z, b + a = z) :=
begin
congrm λ x, ∀ w, _ + a = w,
-- produces one goal for the underscore: ⊢ a = b
exact h,
end
The tactic also allows for "function underscores", denoted by _₁, ..., _₄
. The index denotes
the number of explicit arguments of the function to be matched.
If e
has a "function underscore" in a location, then the tactic reads off the function f
that
appears in lhs
at the current location, replacing the explicit arguments of f
by the user
inputs to the "function underscore". After that, congrm
continues with its matching.