Simp lemmas are used by the "simplifier" family of tactics.
simp_lemmas
is essentially a pair of tables rb_map (expr_type × name) (priority_list simp_lemma)
.
One of the tables is for congruences and one is for everything else.
An individual simp lemma is:
- A kind which can be
Refl
,Simp
orCongr
. - A pair of
expr
sl ~> r
. The rb map is indexed by the name ofget_app_fn(l)
. - A proof that
l = r
orl ↔ r
. - A list of the metavariables that must be filled before the proof can be applied.
- A priority number
Merge the simp_lemma tables.
Remove the given lemmas from the table. Use the names of the lemmas.
Makes the default simp_lemmas table which is composed of all lemmas tagged with simp
.
Add a simplification lemma by an expression p
. Some conditions on p
must hold for it to be added, see list below.
If your lemma is not being added, you can see the reasons by setting set_option trace.simp_lemmas true
.
p
must have the typeΠ (h₁ : _) ... (hₙ : _), LHS ~ RHS
for some reflexive, transitive relation (usually=
).- Any of the hypotheses
hᵢ
should either be present inLHS
or otherwise aProp
or a typeclass instance. LHS
should not occur withinRHS
.LHS
should not occur within a hypothesishᵢ
.
Add a simplification lemma by it's declaration name. See simp_lemmas.add
for more information.
Adds a congruence simp lemma to simp_lemmas.
A congruence simp lemma is a lemma that breaks the simplification down into separate problems.
For example, to simplify a ∧ b
to c ∧ d
, we should try to simp a
to c
and b
to d
.
For examples of congruence simp lemmas look for lemmas with the @[congr]
attribute.
lemma if_simp_congr ... (h_c : b ↔ c) (h_t : x = u) (h_e : y = v) : ite b x y = ite c u v := ...
lemma imp_congr_right (h : a → (b ↔ c)) : (a → b) ↔ (a → c) := ...
lemma and_congr (h₁ : a ↔ c) (h₂ : b ↔ d) : (a ∧ b) ↔ (c ∧ d) := ...
Add expressions to a set of simp lemmas using simp_lemmas.add
.
This is the new version of simp_lemmas.append
,
which also allows you to set the symm
flag.
Add expressions to a set of simp lemmas using simp_lemmas.add
.
This is the backwards-compatibility version of simp_lemmas.append_with_symm
,
and sets all symm
flags to ff
.
simp_lemmas.rewrite s e prove R
apply a simplification lemma from 's'
- 'e' is the expression to be "simplified"
- 'prove' is used to discharge proof obligations.
- 'r' is the equivalence relation being used (e.g., 'eq', 'iff')
'md' is the transparency; how aggresively should the simplifier perform reductions.
Result (new_e, pr) is the new expression 'new_e' and a proof (pr : e R new_e)
simp_lemmas.drewrite s e
tries to rewrite 'e' using only refl lemmas in 's'
(Definitional) Simplify the given expression using only reflexivity equality lemmas from the given set of lemmas. The resulting expression is definitionally equal to the input.
The list u
contains defintions to be delta-reduced, and projections to be reduced.
Tries to unfold e
if it is a constant or a constant application.
Remark: this is not a recursive procedure.
Delta reduce the given constant names
If e
is a projection application, try to unfold it, otherwise fail.
simplify s e cfg r prove
simplify e
using s
using bottom-up traversal.
discharger
is a tactic for dischaging new subgoals created by the simplifier.
If it fails, the simplifier tries to discharge the subgoal by simplifying it to true
.
The parameter to_unfold
specifies definitions that should be delta-reduced,
and projection applications that should be unfolded.
ext_simplify_core a c s discharger pre post r e
:
a : α
- initial user datac : simp_config
- simp configuration optionss : simp_lemmas
- the set of simp_lemmas to use. Remark: the simplification lemmas are not applied automatically like in the simplify tactic. The caller must use them at pre/post.discharger : α → tactic α
- tactic for dischaging hypothesis in conditional rewriting rules. The argument 'α' is the current user data.pre a s r p e
is invoked before visiting the children of subterm 'e'.- arguments:
a
is the current user datas
is the updated set of lemmas if 'contextual' istt
,r
is the simplification relation being used,p
is the "parent" expression (if there is one).e
is the current subexpression in question.- if it succeeds the result is
(new_a, new_e, new_pr, flag)
where new_a
is the new value for the user datanew_e
is a new expression s.t.r e new_e
new_pr
is a proof fore r new_e
, If it is none, the proof is assumed to be by reflexivityflag
if ttnew_e
children should be visited, andpost
invoked.
(post a s r p e)
is invoked after visiting the children of subterme
, The output is similar to(pre a r s p e)
, but the 'flag' indicates whether the new expression should be revisited or not.r
is the simplification relation. Usually=
or↔
.e
is the input expression to be simplified.
The method returns (a,e,pr)
where
a
is the final user datae
is the new expressionpr
is the proof that the given expression equals the input expression.
- to_simp_config : tactic.simp_config
- use_hyps : bool
Example usage:
-- make a new simp attribute called "my_reduction"
run_cmd mk_simp_attr `my_reduction
-- Add "my_reduction" attributes to these if-reductions
attribute [my_reduction] if_pos if_neg dif_pos dif_neg
-- will return the simp_lemmas with the `my_reduction` attribute.
#eval get_user_simp_lemmas `my_reduction
- h : expr
- new_type : expr
- pr : option expr
- s : simp_lemmas