aboutsummaryrefslogtreecommitdiff
path: root/src/Util/Tuple.v
diff options
context:
space:
mode:
authorGravatar Jason Gross <jgross@mit.edu>2017-11-09 21:08:19 -0500
committerGravatar Jason Gross <jgross@mit.edu>2017-11-09 21:08:19 -0500
commit1dff827414482b7fc54fa6764bfd760a0ac93bde (patch)
treec80fda2c4dc2421e1c902e95e6dce3bcb1b97a2b /src/Util/Tuple.v
parente7d670839c8fdf3df92c7cfc4e99a3d5a45dc6b3 (diff)
Generalize fieldwise Proper lemmas
Diffstat (limited to 'src/Util/Tuple.v')
-rw-r--r--src/Util/Tuple.v73
1 files changed, 42 insertions, 31 deletions
diff --git a/src/Util/Tuple.v b/src/Util/Tuple.v
index 8742b3a2b..266bf5f42 100644
--- a/src/Util/Tuple.v
+++ b/src/Util/Tuple.v
@@ -1060,54 +1060,65 @@ Proof.
Qed.
-Global Instance fieldwise'_Proper
- : forall {n A B}, Proper (pointwise_relation _ (pointwise_relation _ impl) ==> eq ==> eq ==> impl) (@fieldwise' A B n) | 10.
+Global Instance fieldwise'_Proper_gen {A B RA RB R}
+ {R_Proper_and : Proper (R ==> R ==> R) and}
+ : forall {n}, Proper ((RA ==> RB ==> R) ==> fieldwise' RA ==> fieldwise' RB ==> R) (@fieldwise' A B n) | 10.
Proof.
induction n as [|n IHn]; intros.
{ compute; intros; subst; auto. }
- { cbv [pointwise_relation Proper respectful impl] in *.
- intros f g Hfg x y ? x' y' ? H'; subst y y'.
- simpl in *; destruct H'; eauto. }
+ { cbv [pointwise_relation Proper respectful] in *.
+ intros f g Hfg x y H x' y' H'.
+ simpl in *; apply R_Proper_and; destruct H, H'; auto. }
Qed.
-Global Instance fieldwise_Proper
- : forall {n A B}, Proper (pointwise_relation _ (pointwise_relation _ impl) ==> eq ==> eq ==> impl) (@fieldwise A B n) | 10.
+Global Instance fieldwise_Proper_gen {A B RA RB R}
+ {R_Proper_and : Proper (R ==> R ==> R) and}
+ (R_True : R True True)
+ : forall {n}, Proper ((RA ==> RB ==> R) ==> fieldwise RA ==> fieldwise RB ==> R) (@fieldwise A B n) | 10.
Proof.
- destruct n; intros; [ | apply fieldwise'_Proper ].
+ destruct n; intros; [ | apply fieldwise'_Proper_gen ].
compute; trivial.
Qed.
-Global Instance fieldwise'_Proper_iff
- : forall {n A B}, Proper (pointwise_relation _ (pointwise_relation _ iff) ==> eq ==> eq ==> iff) (@fieldwise' A B n) | 10.
+Global Instance fieldwise'_Proper_gen_eq {A B R}
+ {R_Proper_and : Proper (R ==> R ==> R) and}
+ : forall {n}, Proper (pointwise_relation _ (pointwise_relation _ R) ==> eq ==> eq ==> R) (@fieldwise' A B n) | 10.
Proof.
- induction n as [|n IHn]; intros.
- { compute; intros; subst; auto. }
- { cbv [pointwise_relation Proper respectful] in *.
- intros f g Hfg x y ? x' y' ?; subst y y'.
- simpl; erewrite IHn by first [ eassumption | eauto ].
- rewrite Hfg; reflexivity. }
+ cbv [pointwise_relation].
+ repeat intro; subst; apply (@fieldwise'_Proper_gen A B eq eq); try assumption;
+ repeat intro; subst; auto; reflexivity.
Qed.
-Global Instance fieldwise_Proper_iff
- : forall {n A B}, Proper (pointwise_relation _ (pointwise_relation _ iff) ==> eq ==> eq ==> iff) (@fieldwise A B n) | 10.
+Global Instance fieldwise_Proper_gen_eq {A B R}
+ {R_Proper_and : Proper (R ==> R ==> R) and}
+ (R_True : R True True)
+ : forall {n}, Proper (pointwise_relation _ (pointwise_relation _ R) ==> eq ==> eq ==> R) (@fieldwise A B n) | 10.
Proof.
- destruct n; intros; [ | apply fieldwise'_Proper_iff ].
- compute; tauto.
+ cbv [pointwise_relation].
+ repeat intro; subst; apply (@fieldwise_Proper_gen A B eq eq); try assumption;
+ repeat intro; subst; auto; reflexivity.
Qed.
+Global Instance fieldwise'_Proper
+ : forall {n A B}, Proper (pointwise_relation _ (pointwise_relation _ impl) ==> eq ==> eq ==> impl) (@fieldwise' A B n) | 10.
+Proof. intros; apply fieldwise'_Proper_gen_eq. Qed.
+
+Global Instance fieldwise_Proper
+ : forall {n A B}, Proper (pointwise_relation _ (pointwise_relation _ impl) ==> eq ==> eq ==> impl) (@fieldwise A B n) | 10.
+Proof. intros; now apply fieldwise_Proper_gen_eq. Qed.
+
+Global Instance fieldwise'_Proper_iff
+ : forall {n A B}, Proper (pointwise_relation _ (pointwise_relation _ iff) ==> eq ==> eq ==> iff) (@fieldwise' A B n) | 10.
+Proof. intros; apply fieldwise'_Proper_gen_eq. Qed.
+
+Global Instance fieldwise_Proper_iff
+ : forall {n A B}, Proper (pointwise_relation _ (pointwise_relation _ iff) ==> eq ==> eq ==> iff) (@fieldwise A B n) | 10.
+Proof. intros; now apply fieldwise_Proper_gen_eq. Qed.
+
Global Instance fieldwise'_Proper_flip_impl
: forall {n A B}, Proper (pointwise_relation _ (pointwise_relation _ (flip impl)) ==> eq ==> eq ==> flip impl) (@fieldwise' A B n) | 10.
-Proof.
- induction n as [|n IHn]; intros.
- { compute; intros; subst; auto. }
- { cbv [pointwise_relation Proper respectful flip impl] in *.
- intros f g Hfg x y ? x' y' ? H'; subst y y'.
- simpl in *; destruct H'; eauto. }
-Qed.
+Proof. intros; apply @fieldwise'_Proper_gen_eq; compute; tauto. Qed.
Global Instance fieldwise_Proper_flip_impl
: forall {n A B}, Proper (pointwise_relation _ (pointwise_relation _ (flip impl)) ==> eq ==> eq ==> flip impl) (@fieldwise A B n) | 10.
-Proof.
- destruct n; intros; [ | apply fieldwise'_Proper_flip_impl ].
- compute; trivial.
-Qed.
+Proof. intros; apply @fieldwise_Proper_gen_eq; compute; tauto. Qed.