Library coqrel.Monotonicity
Require Export RelDefinitions.
Require Export RelOperators.
Require Export Relators.
Require Import Delay.
Require Export RelOperators.
Require Export Relators.
Require Import Delay.
The monotonicity tactic
Truncating applications
If the head term on either side is an existential variable, we
only check Params declarations for the other side. This auxiliary
lemma is used to trigger this step; p should be the number of
unapplied arguments in whichever term m1 of m2 we extracted the
non-existential head term h from.
Lemma query_params_one {A B C} (h: A) p (m1: B) (m2: C) n:
Params h (p + n) →
QueryParams m1 m2 n.
Proof.
constructor.
Qed.
If neither head term is an existential variable, then the Params
declarations for both should agree.
Lemma query_params_both {A B C D} (h1: A) (h2: B) p1 p2 (m1: C) (m2: D) n:
Params h1 (p1 + n) →
Params h2 (p2 + n) →
QueryParams m1 m2 n.
Proof.
constructor.
Qed.
Now we need to choose which one of these lemmas to apply in a
given situation, and compute the appropriate values for p.
We compute p by essentially counting the products in the type of
the related terms m. Note that we need to make sure we "open up"
types such as subrel to expose the products, accomplished here by
using eval cbv.
FIXME: count_unapplied_params fails if m has a dependent type,
due to t' being a term under binders rather than a closed term.
This means monotonicity can't handle goal where the related terms
have dependent types if the corresponding relational property is
about a partial application.
Ltac count_unapplied_params m :=
let rec count t :=
lazymatch t with
| ∀ x, ?t' ⇒ let n := count t' in constr:(S n)
| _ ⇒ constr:(O)
end in
let t := type of m in
let t := eval cbv in t in
count t.
Ltac query_params m1 m2 :=
let rec head t := lazymatch t with ?t' _ ⇒ head t' | _ ⇒ t end in
let h1 := head m1 in
let p1 := count_unapplied_params m1 in
let h2 := head m2 in
let p2 := count_unapplied_params m2 in
first
[ not_evar h1; not_evar h2; eapply (query_params_both h1 h2 p1 p2)
| not_evar h1; is_evar h2; eapply (query_params_one h1 p1)
| is_evar h1; not_evar h2; eapply (query_params_one h2 p2) ].
Hint Extern 10 (QueryParams ?m1 ?m2 _) ⇒
query_params m1 m2 : typeclass_instances.
Now that we can figure out how many parameters to drop, we use
that information to construct the pair of prefixes we use to locate
the relational property. We try the following (in this order):
Again, we need to take into account that either application could
contain an eixstential variable. When we encounter an existential
variable while chopping off arguments, we short-circuit the process
and simply generate a new evar to serve as the shortened version.
- the whole applications f x1 ... xn and g y1 ... yn;
- prefixes constructed as above from user-declared Params instances;
- the smallest prefix we can obtain by dropping arguments from both applications simulataneously.
Ltac pass_evar_to k :=
let Av := fresh "A" in evar (Av : Type);
let A := eval red in Av in clear Av;
let av := fresh "a" in evar (av : A);
let a := eval red in av in clear av;
k a.
The class RemoveParams implements the concurrent removal of n
parameters from applications m1 and m2.
Class RemoveParams {A B C D} (n: nat) (m1: A) (m2: B) (f: C) (g: D).
Ltac remove_params_from n m k :=
lazymatch n with
| O ⇒ k m
| S ?n' ⇒
lazymatch m with
| ?m' _ ⇒ remove_params_from n' m' k
| _ ⇒ is_evar m; pass_evar_to k
end
end.
Ltac remove_params n m1 m2 f g :=
remove_params_from n m1 ltac:(fun m1' ⇒ unify f m1';
remove_params_from n m2 ltac:(fun m2' ⇒ unify g m2')).
Hint Extern 1 (RemoveParams ?n ?m1 ?m2 ?f ?g) ⇒
remove_params n m1 m2 f g; constructor : typeclass_instances.
The class RemoveAllParams implements the concurrent removal of
as many arguments as possible.
Class RemoveAllParams {A B C D} (m1: A) (m2: B) (f: C) (g: D).
Ltac remove_all_params_from m k :=
lazymatch m with
| ?m' _ ⇒ remove_all_params_from m' k
| _ ⇒ not_evar m; k m
end.
Ltac remove_all_params m1 m2 f g :=
first
[ is_evar m2;
remove_all_params_from m1
ltac:(fun m1' ⇒ unify f m1'; pass_evar_to ltac:(fun m2' ⇒ unify g m2'))
| is_evar m1;
remove_all_params_from m2
ltac:(fun m2' ⇒ unify g m2'; pass_evar_to ltac:(fun m1' ⇒ unify f m1'))
| lazymatch m1 with ?m1' _ ⇒
lazymatch m2 with ?m2' _ ⇒
remove_all_params m1' m2' f g
end
end
| not_evar m1; unify f m1;
not_evar m2; unify g m2 ].
Hint Extern 1 (RemoveAllParams ?m1 ?m2 ?f ?g) ⇒
remove_all_params m1 m2 f g; constructor : typeclass_instances.
Selecting a relational property
Class CandidateProperty {A B} (R': rel A B) f g (Q: Prop) :=
candidate_related: R' f g.
Instance as_is_candidate {A B} (R R': rel A B) m1 m2:
Related m1 m2 R' →
CandidateProperty R' m1 m2 (R m1 m2) | 10.
Proof.
firstorder.
Qed.
Instance remove_params_candidate {A B C D} R (m1: A) (m2: B) R' (f: C) (g: D) n:
QueryParams m1 m2 n →
RemoveParams n m1 m2 f g →
Related f g R' →
CandidateProperty R' f g (R m1 m2) | 20.
Proof.
firstorder.
Qed.
Instance remove_all_params_candidate {A B C D} R (m1:A) (m2:B) R' (f:C) (g:D):
RemoveAllParams m1 m2 f g →
Related f g R' →
CandidateProperty R' f g (R m1 m2) | 30.
Proof.
firstorder.
Qed.
We also attempt to use any matching hypothesis. This takes
priority and bypasses any Params declaration. We assume that we
will always want such hypotheses to be used, at least when the left-
or right-hand side matches exactly. There is a possibility that this
ends up being too broad for some applications, for which we'll want
to restrict ourselves to Related instances explicitely defined by
the user. If this turns out to be the case, we can introduce an
intermediate class with more parameters to control the sources of
the relational properties we use, and perhaps have some kind of
normalization process akin to what is done in Coq.Classes.Morphisms.
We don't hesitate to instantiate evars in the goal or hypothesis; if
the types or skeletons are compatible this is likely what we want.
In one practical case where this is needed, my hypothesis is
something like R (f ?x1 ?y1 42) (f ?x2 ?y2 24), and I want
?x1, ?x2, ?y1, ?y2 to be matched against the goal. In case
the use of unify below is too broad, we could develop a strategy
whereby at least one of the head terms f, g should match exactly.
We also consider the possibility that the relation's arguments may
be flipped in the hypothesis, compared to the goal, when the
relation is of the form (or can be instantiated to) flip _.
In that case flip_relim will allow us to use the property.
This is especially common in the course of solving a goal of the
form (R --> R') f g, or goals generated by the setoid rewriting
system of the form Proper (?R1 ==> ... ==> ?Rn ==> flip impl),
and where the ?Ris will generally need to be instantiated as
flip ?Ri'. We do not pursue a similar strategy when looking up
candidates from the typeclass instance database, because we expect
the user to normalize such properties before they declare them, for
example declare an instance of Related (R1 ++> R2 --> R3) g f
rather than an instance of Related (R1 --> R2 ++> flip R3) f g.
Note that it is important that we reduce the goal to ?R ?m ?n
before we use eexact: if the relation in the hypothesis is an
existential variable, we don't want unified against
CandidateProperty _ _ _ (_ ?m ?n).
Lemma unflip_context_candidate {A B} (R: rel A B) x y : flip R y x → R x y.
Proof.
firstorder.
Qed.
Ltac context_candidate :=
let rec is_prefix f m :=
first
[ unify f m
| lazymatch m with ?n _ ⇒ is_prefix f n end ] in
let rec is_prefixable f m :=
first
[ is_evar f
| is_evar m
| unify f m
| lazymatch m with ?n _ ⇒ is_prefixable f n end ] in
multimatch goal with
| H: _ ?f ?g |- @CandidateProperty ?A ?B ?R ?x ?y (_ ?m ?n) ⇒
red;
once first
[ is_prefix f m; is_prefixable g n
| is_prefix g n; is_prefixable f m
| is_prefix g m; is_prefixable f n;
apply (@unflip_context_candidate A B R x y)
| is_prefix f n; is_prefixable g m;
apply (@unflip_context_candidate A B R x y)];
eexact H
end.
Hint Extern 1 (CandidateProperty _ _ _ _) ⇒
context_candidate : typeclass_instances.
Using subrel
While we give priority to rimpl_refl below, when it doesn't work
we use RAuto to establish a subrel property. The Monotonic
instances we declared alongside relators can be used conjunction
with Monotonicity to break up subrel goals along the structure
of the relations being considered. This may end up causing loops
(especially in failure cases), so it may be necessary to add some
flag in the context to prevent subrel_related from being used when
discharging the subrel goals themselves.
Global Instance rimpl_refl {A B} (R : rel A B) m n:
RImpl (R m n) (R m n).
Proof.
firstorder.
Qed.
Global Instance rimpl_subrel {A B} (R R': rel A B) m n:
NonDelayed (RAuto (subrel R R')) →
RImpl (R m n) (R' m n).
Proof.
firstorder.
Qed.
Main tactic
Class Monotonicity (P Q: Prop): Prop :=
monotonicity: P → Q.
Global Instance apply_candidate {A B} (R: rel A B) m n P Q Q':
CandidateProperty R m n Q →
RElim R m n P Q' →
RImpl Q' Q →
Monotonicity P Q.
Proof.
firstorder.
Qed.
The Ltac tactic simply applies monotonicity; typeclass
resolution will do the rest. Note that using apply naively is too
lenient because in a goal of type A → B, it will end up unifying
A with P and B with Q instead of setting Q := A → B and
generating a new subgoal for P as expected. On the other hand,
using refine directly is too restrictive because it will not unify
the type of monotonicity against the goal if existential variables
are present in one or the other. Hence we constrain apply just
enough, so as to handle both of these cases.
Ltac monotonicity :=
lazymatch goal with |- ?Q ⇒ apply (monotonicity (Q:=Q)) end;
Delay.split_conjunction.
Another way to use Monotonicity is to hook it as an RStep
instance.
Global Instance monotonicity_rstep {A B} (P: Prop) (R: rel A B) m n:
Monotonicity P (R m n) →
RStep P (R m n) | 50.
Proof.
firstorder.
Qed.
subrel constraint as a subgoal
Class SubrelMonotonicity (P Q: Prop): Prop :=
subrel_monotonicity: P → Q.
Global Instance apply_candidate_subrel {A B C D} (R: rel A B) m n P Rc Rg x y:
CandidateProperty R m n (Rg x y) →
RElim R m n P (Rc x y) →
SubrelMonotonicity (@subrel C D Rc Rg ∧ P) (Rg x y).
Proof.
firstorder.
Qed.
Ltac subrel_monotonicity :=
lazymatch goal with |- ?Q ⇒ apply (subrel_monotonicity (Q:=Q)) end;
Delay.split_conjunction.
Generic instances
Class CommonPrefix {A B C} (m1: A) (m2: B) (f: C).
Ltac common_prefix m1 m2 f :=
first
[ not_evar m1; not_evar m2;
unify m1 m2; unify f m1
| lazymatch m1 with ?m1' _ ⇒
lazymatch m2 with ?m2' _ ⇒
common_prefix m1' m2' f
end
end
| unify m1 m2; unify f m1 ].
Hint Extern 1 (CommonPrefix ?m1 ?m2 ?f) ⇒
common_prefix m1 m2 f; constructor : typeclass_instances.
Global Instance eq_candidate {A B C} R (m1: A) (m2: B) (f: C):
CommonPrefix m1 m2 f →
CandidateProperty eq f f (R m1 m2) | 100.
Proof.
firstorder.
Qed.
Then, we can use the following RElim instance to obtain the
behavior of the f_equal tactic.
In practice, I find that f_equal_relim is too broadly applicable.
In situations where a relational property is missing or inapplicable
for some reason, we would like repeat rstep to leave us in a
situation where we can examine what is going wrong, or do a manual
proof when necessary. If we use f_equal-like behavior as a
default, we will instead be left several step further, where it is
no longer obvious which function is involved, and where the goal may
not be provable anymore.
With that said, if you would like to switch f_equal on, you can
register f_equal_elim as an instance with the following hint:
Hint Extern 1 (RElim (@eq (_ -> _)) _ _ _ _) => eapply f_equal_relim : typeclass_instances.
Lemma f_equal_relim {A B} f g m n P Q:
RElim eq (f m) (g n) P Q →
RElim (@eq (A → B)) f g (m = n ∧ P) Q.
Proof.
intros Helim Hf [Hmn HP].
apply Helim; eauto.
congruence.
Qed.
Note that thanks to eq_subrel, this will also apply to a goal
that uses a Reflexive relation, not just a goal that uses eq.
In fact, this subsumes the reflexivity tactic as well, which
corresponds to the special case where all arguments are equal;
in that case relim_base can be used directly and f_equal_elim is
not necessary.