├── .gitignore ├── BoolRel.v ├── CITATION.cff ├── Delay.v ├── KLR.v ├── LICENSE ├── LogicalRelations.v ├── LogicalRelationsTests.v ├── Monotonicity.v ├── MorphismsInterop.v ├── OptionRel.v ├── PreOrderTactic.v ├── RDestruct.v ├── RelClasses.v ├── RelDefinitions.v ├── RelOperators.v ├── Relators.v ├── Transport.v ├── _CoqProject ├── configure └── index.html /.gitignore: -------------------------------------------------------------------------------- 1 | Makefile 2 | Makefile.bak 3 | Makefile.conf 4 | .coqdeps.d 5 | .Makefile.d 6 | *.v.d 7 | *.vo 8 | *.vok 9 | *.vos 10 | *.glob 11 | .*.aux 12 | -------------------------------------------------------------------------------- /BoolRel.v: -------------------------------------------------------------------------------- 1 | Require Import LogicalRelations. 2 | Require Export Coq.Bool.Bool. 3 | 4 | 5 | (** * Relations for [sumbool] *) 6 | 7 | (** ** Symmetric relation *) 8 | 9 | (** NB: With proof irrelevance this is the same as [eq]. *) 10 | 11 | Inductive sumbool_rel {P1 P2 Q1 Q2}: rel (sumbool P1 Q1) (sumbool P2 Q2):= 12 | | left_rel_def H1 H2: sumbool_rel (left H1) (left H2) 13 | | right_rel_def H1 H2: sumbool_rel (right H1) (right H2). 14 | 15 | Global Instance left_rel: 16 | Monotonic (@left) (forallr _ _ : ⊤, forallr _ _ : ⊤, ⊤ ++> sumbool_rel). 17 | Proof. 18 | constructor; eauto. 19 | Qed. 20 | 21 | Global Instance right_rel: 22 | Monotonic (@right) (forallr _ _ : ⊤, forallr _ _ : ⊤, ⊤ ++> sumbool_rel). 23 | Proof. 24 | constructor; eauto. 25 | Qed. 26 | 27 | (** ** Directed relation *) 28 | 29 | Inductive sumbool_le {P1 P2 Q1 Q2}: rel (sumbool P1 Q1) (sumbool P2 Q2):= 30 | | left_le b1 H2: Related b1 (left H2) sumbool_le 31 | | right_le H1 b2: Related (right H1) b2 sumbool_le. 32 | 33 | Global Existing Instance left_le. 34 | Global Existing Instance right_le. 35 | 36 | 37 | (** * Definitions for [Bool.le] *) 38 | 39 | (** ** Properties *) 40 | 41 | Global Instance leb_preo: 42 | PreOrder Bool.le. 43 | Proof. 44 | split. 45 | - intros [|]; simpl; eauto. 46 | - intros [|]; simpl; eauto. 47 | intros [|]; simpl; eauto. 48 | discriminate. 49 | Qed. 50 | 51 | Lemma leb_rdestruct: 52 | RDestruct Bool.le (fun P => P false false /\ P true true /\ P false true). 53 | Proof. 54 | intros b1 b2 Hb P (HPff & HPtt & HPft). 55 | destruct b1, b2; eauto; discriminate. 56 | Qed. 57 | 58 | Global Hint Extern 0 (RDestruct Bool.le _) => 59 | eapply leb_rdestruct : typeclass_instances. 60 | 61 | Global Instance leb_transport_eq_true x y: 62 | Transport Bool.le x y (x = true) (y = true). 63 | Proof. 64 | clear. 65 | destruct x, y; firstorder. 66 | Qed. 67 | 68 | (** ** Monotonicity of various definitions *) 69 | 70 | Lemma true_leb b: 71 | Related b true Bool.le. 72 | Proof. 73 | destruct b; reflexivity. 74 | Qed. 75 | 76 | Global Hint Extern 0 (Related _ true _) => 77 | eapply true_leb : typeclass_instances. 78 | 79 | Lemma false_leb b: 80 | Related false b Bool.le. 81 | Proof. 82 | destruct b; reflexivity. 83 | Qed. 84 | 85 | Global Hint Extern 0 (Related false _ _) => 86 | eapply false_leb : typeclass_instances. 87 | 88 | Global Instance negb_leb: 89 | Monotonic negb (Bool.le --> Bool.le). 90 | Proof. 91 | unfold negb. 92 | rauto. 93 | Qed. 94 | 95 | Global Instance andb_leb: 96 | Monotonic andb (Bool.le ++> Bool.le ++> Bool.le). 97 | Proof. 98 | unfold andb. 99 | rauto. 100 | Qed. 101 | 102 | Global Instance orb_leb: 103 | Monotonic orb (Bool.le ++> Bool.le ++> Bool.le). 104 | Proof. 105 | unfold orb. 106 | rauto. 107 | Qed. 108 | 109 | Global Instance implb_leb: 110 | Monotonic implb (Bool.le --> Bool.le ++> Bool.le). 111 | Proof. 112 | unfold implb. 113 | rauto. 114 | Qed. 115 | -------------------------------------------------------------------------------- /CITATION.cff: -------------------------------------------------------------------------------- 1 | # This CITATION.cff file was generated with cffinit. 2 | # Visit https://bit.ly/cffinit to generate yours today! 3 | 4 | cff-version: 1.2.0 5 | title: >- 6 | Coqrel: a binary logical relations library for the Coq 7 | proof assistant 8 | message: >- 9 | If you use this software, please cite it using the 10 | metadata from this file. 11 | type: software 12 | authors: 13 | - given-names: Jérémie 14 | family-names: Koenig 15 | email: jeremie.koenig@yale.edu 16 | affiliation: Yale University 17 | orcid: 'https://orcid.org/0000-0002-3168-5925' 18 | repository-code: 'https://github.com/CertiKOS/coqrel' 19 | url: 'http://certikos.github.io/coqrel/' 20 | license: MIT 21 | -------------------------------------------------------------------------------- /Delay.v: -------------------------------------------------------------------------------- 1 | (** This file implements a tactical [delayed tac] and a tactic 2 | [delay], which function in the following way: [tac] is able to 3 | invoke [delay] to solve any subgoal. Assuming the subgoal can be 4 | typed in the context in which [delayed tac] was invoked, the 5 | [delay] tactic will succeed, and the subgoal will instead become a 6 | subgoal of the encolosing invocation of [delayed tac]. 7 | 8 | To see why this is useful, consider the way [eauto] works. Applying 9 | an [eauto] hint is only considered successful when the generated 10 | subgoal can themselves be solved by [eauto]. Otherwise, [eauto] 11 | backtracks and another hint is used. If no hint can be used to solve 12 | the goal completely, then no progress is made. 13 | 14 | While this behavior is usually appropriate, this means [eauto] 15 | cannot be used as a base for tactics which make as much progress as 16 | possible, then let the user (or another tactic in a sequence) solve 17 | the remaining subgoals. With the tactics defined below, this can be 18 | accomplished by registering an external hint making use of [delay] 19 | for strategic subgoals, and invoking [delayed eauto]. 20 | 21 | Additionally, the variant [delayed_conjunction] bundles all of the 22 | subgoals solved by [delay] in a single conjunction of the form 23 | [P1 /\ ... /\ Pn /\ True]. This is useful as it provides a uniform 24 | interface when reifying the effect of a tactic. *) 25 | 26 | 27 | (** * The [delayed]/[delay] tactics *) 28 | 29 | Module Delay. 30 | (** The [delayed] tactical and the [delay] tactic communicate 31 | through the use of existential variables. Specifically, [delayed] 32 | introduces a so-called "open conjunction" to the context. The open 33 | conjunction starts as a hypothesis of type [?P], and remains of 34 | the form [P1 /\ ... /\ Pn /\ ?Q] throughout. Each invocation of 35 | [delay] unifies [?Q] with a new node [Pi /\ ?Q']. [Pi] will be of 36 | the form [forall x1 .. xn, delayed_goal G], where [G] is the 37 | original goal and the [xi]s are the new variables which were 38 | introduced after the open conjunction was created. *) 39 | 40 | Ltac use_conjunction H := 41 | lazymatch type of H with 42 | | ?A /\ ?B => 43 | use_conjunction (@proj2 A B H) 44 | | _ => 45 | eapply (proj1 H) 46 | end. 47 | 48 | (** As much as possible, our open conjunction should remain hidden 49 | from the user's point of view. To avoid interfering with random 50 | tactics such as [eassumption], we wrap it in the following record 51 | type. *) 52 | 53 | Record open_conjunction (P: Prop) := 54 | { 55 | open_conjunction_proj: P 56 | }. 57 | 58 | (** The evar can only be instantiated to refer to variables that 59 | existed at the point where the evar is spawned. As a consequence, 60 | in order to use it, we must first get rid of all of the new 61 | variables, which the goal potentially depends on. We do this by 62 | reverting the context one variable at a time from the bottom up, 63 | until we hit our [open_conjunction]. *) 64 | 65 | Ltac revert_until_conjunction Hdelayed := 66 | match goal with 67 | | |- open_conjunction _ -> _ => 68 | intro Hdelayed 69 | | H : _ |- _ => 70 | revert H; 71 | revert_until_conjunction Hdelayed 72 | end. 73 | 74 | (** To make sure this operation is reversible, we first wrap the 75 | goal using the following marker. *) 76 | 77 | Definition delayed_goal (P: Prop) := P. 78 | 79 | (** In some contexts, we actually want a nested saved goal to be 80 | split up further instead. The following marker does that. *) 81 | 82 | Definition unpack (P: Prop) := P. 83 | 84 | (** Once we're done, we've packaged all of the goals we solved using 85 | [use_conjunction] into a single hypothesis. We can use the 86 | following tactic to split up that conjunction into individual 87 | subgoals again. 88 | 89 | Although we expect the conjunctions produced by [use_conjunction] 90 | to be terminated by [True] and this should be regarded as the 91 | canonical format, when the last conjunct is a non-[True] 92 | proposition we consider it as an additional subgoal. Likewise we 93 | don't require the conclusions to be wrapped in [delayed_goal]. 94 | This allows for some flexibility when dealing with manually 95 | written goals. *) 96 | 97 | Ltac split_conjunction := 98 | let handle_subgoal := 99 | intros; 100 | match goal with 101 | | |- delayed_goal _ => red 102 | | |- unpack _ => red; split_conjunction 103 | | |- _ => idtac 104 | end in 105 | match goal with 106 | | |- _ /\ _ => split; [handle_subgoal | split_conjunction] 107 | | |- True => exact I 108 | | |- ?P => is_evar P; exact I 109 | | |- _ => handle_subgoal 110 | end. 111 | 112 | (** Now we can defined [delayed], [delayed_conjunction] and [delay]. *) 113 | 114 | Ltac delay := 115 | idtac; 116 | lazymatch goal with 117 | | _ : open_conjunction ?P |- ?G => 118 | change (delayed_goal G); 119 | let Hdelayed := fresh "Hdelayed" in 120 | revert_until_conjunction Hdelayed; 121 | use_conjunction (open_conjunction_proj P Hdelayed) 122 | | _ => 123 | fail "delay can only be used under the 'delayed' tactical" 124 | end. 125 | 126 | Ltac delayed_conjunction tac := 127 | let Pv := fresh in evar (Pv: Prop); 128 | let P := eval red in Pv in clear Pv; 129 | let Hdelayed := fresh "Hdelayed" in 130 | cut (open_conjunction P); 131 | [ intro Hdelayed; tac; clear Hdelayed | 132 | eapply Build_open_conjunction ]. 133 | 134 | Ltac delayed tac := 135 | delayed_conjunction tac; 136 | [ .. | split_conjunction ]. 137 | 138 | Ltac undelay := 139 | lazymatch goal with 140 | | Hdelayed : open_conjunction _ |- _ => 141 | clear Hdelayed 142 | end. 143 | 144 | Ltac nondelayed tac := 145 | undelay; tac. 146 | End Delay. 147 | 148 | Tactic Notation "delayed" tactic1(tac) := 149 | Delay.delayed tac. 150 | 151 | Tactic Notation "nondelayed" tactic1(tac) := 152 | Delay.nondelayed tac. 153 | 154 | Tactic Notation "delayed_conjunction" tactic1(tac) := 155 | Delay.delayed_conjunction tac. 156 | 157 | Tactic Notation "delay" := Delay.delay. 158 | 159 | 160 | (** * Hints *) 161 | 162 | (** We hook [delay] into [eauto] in two different ways: subgoals 163 | intended to be delayed can be marked with the wrapper below. 164 | Alternately, one can use the [delay] hint database, which provides a 165 | low-priority delay rule. *) 166 | 167 | Definition delay (P: Prop) := P. 168 | 169 | Global Hint Extern 0 (delay _) => delay : core. 170 | 171 | Global Hint Extern 100 => delay : delay. 172 | 173 | (** This typeclass wrapper is convenient for performing a nested 174 | search outside of a potential "delayed" context. *) 175 | 176 | Class NonDelayed (P: Prop) := 177 | nondelayed : P. 178 | 179 | Global Hint Extern 1 (NonDelayed _) => 180 | red; try Delay.undelay : typeclass_instances. 181 | 182 | 183 | (** * Example: reifying [eapply] *) 184 | 185 | (** To demonstrate how [delayed_conjunction] can be used to reify a 186 | tactic into a typeclass, we define the [EApply] typeclass. An 187 | instance of [EApply HT P Q] expresses that if [H] is a hypothesis of 188 | type [HT], then invoking [eapply H] on a goal of type [Q] will 189 | generate the subgoal conjunction [P]. *) 190 | 191 | Class EApply (HT P Q: Prop) := 192 | eapply : HT -> P -> Q. 193 | 194 | Global Hint Extern 1 (EApply ?HT _ _) => 195 | let H := fresh in 196 | let HP := fresh "HP" in 197 | intros H HP; 198 | delayed_conjunction solve [eapply H; delay]; 199 | eexact HP : 200 | typeclass_instances. 201 | -------------------------------------------------------------------------------- /KLR.v: -------------------------------------------------------------------------------- 1 | Require Export Monotonicity. 2 | 3 | (** Kripke logical relations generalize logical relations along the 4 | lines of the Kripke semantics of modal logic. A Kripke logical 5 | relation is indexed by a set of worlds equipped with an 6 | accessibility relation. 7 | 8 | This is useful when reasoning about stateful systems, because in 9 | that context the way we want to relate different components may 10 | evolve with the system state. Then, the Kripke world is essentially 11 | the relational counterpart to the state of the system, and the 12 | accessibility relation specifies how this world may evolve alongside 13 | the state (it denotes a notion of "possible future", for instance). *) 14 | 15 | 16 | (** * Kripke frames and logical relations *) 17 | 18 | (** Kripke frames specify the set of worlds over which a KLR is 19 | indexed and an associated accessibility relation. For a given Kripke 20 | frame, the type [klr] is a [W]-indexed version of [rel]. *) 21 | 22 | Definition klr W A B: Type := 23 | W -> rel A B. 24 | 25 | Class KripkeFrame (L: Type) (W: Type) := 26 | { 27 | acc: L -> rel W W; 28 | }. 29 | 30 | Infix "~>" := (acc tt) (at level 70). 31 | 32 | Declare Scope klr_scope. 33 | Delimit Scope klr_scope with klr. 34 | Bind Scope klr_scope with klr. 35 | 36 | 37 | (** * Operations on KLRs *) 38 | 39 | Inductive klr_sum {WA A1 A2 WB B1 B2} RA RB: klr (WA + WB) (A1 + B1) (A2 + B2) := 40 | | klr_inl w a1 a2 : 41 | RA w a1 a2 -> 42 | klr_sum RA RB (inl w) (inl a1) (inl a2) 43 | | klr_inr w b1 b2 : 44 | RB w b1 b2 -> 45 | klr_sum RA RB (inr w) (inr b1) (inr b2). 46 | 47 | 48 | (** * Kripke relators *) 49 | 50 | (** Just like relators allow us to construct complex relations from 51 | simpler ones in a structure-preserving way, Kripke relators allow us 52 | to build complex Kripke relations from simpler Kripke relations. *) 53 | 54 | (** ** Lifted relators *) 55 | 56 | (** First, we lift usual relators to obtain a Kripke version, 57 | basically as the pointwise extension over worlds of the original 58 | ones. *) 59 | 60 | (** *** Lifiting operators *) 61 | 62 | Section LIFT. 63 | Context `{kf: KripkeFrame}. 64 | Context {A1 B1 A2 B2 A B: Type}. 65 | 66 | (** For elementary relations, the corresponding Kripke relation can 67 | just ignore the Kripke world. *) 68 | 69 | Definition k (R: rel A B): klr W A B := 70 | fun w => R. 71 | 72 | Global Instance k_rel var: 73 | Monotonic k (var ++> - ==> var). 74 | Proof. 75 | unfold k; rauto. 76 | Qed. 77 | 78 | Lemma k_rintro R (w: W) (x: A) (y: B): 79 | RIntro (R x y) (k R w) x y. 80 | Proof. 81 | firstorder. 82 | Qed. 83 | 84 | Lemma k_relim (R: rel A B) w x y P Q: 85 | RElim R x y P Q -> 86 | RElim (k R w) x y P Q. 87 | Proof. 88 | tauto. 89 | Qed. 90 | 91 | (** For relators of higher arities, we let the original relator [RR] 92 | operate independently at each world [w]. *) 93 | 94 | Definition k1 RR (R1: klr W A1 B1): klr W A B := 95 | fun w => RR (R1 w). 96 | 97 | Global Instance k1_rel var1 var: 98 | Monotonic k1 ((var1 ++> var) ++> ((- ==> var1) ++> (- ==> var))). 99 | Proof. 100 | unfold k1; rauto. 101 | Qed. 102 | 103 | Lemma k1_rintro RR (R1: klr W A1 B1) w x y: 104 | RIntro (RR (R1 w) x y) (k1 RR R1 w) x y. 105 | Proof. 106 | firstorder. 107 | Qed. 108 | 109 | Lemma k1_relim RR (R1: klr W A1 B1) w x y P Q: 110 | RElim (RR (R1 w)) x y P Q -> 111 | RElim (k1 RR R1 w) x y P Q. 112 | Proof. 113 | tauto. 114 | Qed. 115 | 116 | Definition k2 RR (R1: klr W A1 B1) (R2: klr W A2 B2): klr W A B := 117 | fun w => RR (R1 w) (R2 w). 118 | 119 | Global Instance k2_rel var1 var2 var: 120 | Monotonic k2 121 | ((var1 ++> var2 ++> var) ++> 122 | ((- ==> var1) ++> (- ==> var2) ++> (- ==> var))). 123 | Proof. 124 | unfold k2; rauto. 125 | Qed. 126 | 127 | Lemma k2_rintro RR (R1: klr W A1 B1) (R2: klr W A2 B2) w x y: 128 | RIntro (RR (R1 w) (R2 w) x y) (k2 RR R1 R2 w) x y. 129 | Proof. 130 | firstorder. 131 | Qed. 132 | 133 | Lemma k2_relim RR (R1: klr W A1 B1) (R2: klr W A2 B2) w x y P Q: 134 | RElim (RR (R1 w) (R2 w)) x y P Q -> 135 | RElim (k2 RR R1 R2 w) x y P Q. 136 | Proof. 137 | tauto. 138 | Qed. 139 | 140 | Global Instance k2_unfold RR (R1: klr W A1 B1) (R2: klr W A2 B2) w: 141 | Related (RR (R1 w) (R2 w)) (k2 RR R1 R2 w) subrel. 142 | Proof. 143 | red. reflexivity. 144 | Qed. 145 | End LIFT. 146 | 147 | Global Instance k_rel_params: Params (@k) 4 := { }. 148 | Global Instance k1_rel_params: Params (@k1) 5 := { }. 149 | Global Instance k2_rel_params: Params (@k2) 6 := { }. 150 | 151 | Global Hint Extern 0 (RIntro _ (k _ _) _ _) => 152 | eapply k_rintro : typeclass_instances. 153 | Global Hint Extern 1 (RElim (k _ _) _ _ _ _) => 154 | eapply k_relim : typeclass_instances. 155 | 156 | Global Hint Extern 0 (RIntro _ (k1 _ _ _) _ _) => 157 | eapply k1_rintro : typeclass_instances. 158 | Global Hint Extern 1 (RElim (k1 _ _ _) _ _ _ _) => 159 | eapply k1_relim : typeclass_instances. 160 | 161 | Global Hint Extern 0 (RIntro _ (k2 _ _ _ _) _ _) => 162 | eapply k2_rintro : typeclass_instances. 163 | Global Hint Extern 1 (RElim (k2 _ _ _ _) _ _ _ _) => 164 | eapply k2_relim : typeclass_instances. 165 | 166 | (** *** Usual relators *) 167 | 168 | (** Using the lifting operators defined above, we can provide a set of 169 | usual Kripke relators for various types. *) 170 | 171 | Section USUAL. 172 | Context `{kf: KripkeFrame}. 173 | 174 | Definition arrow_klr {A1 A2 B1 B2} := 175 | k2 (W:=W) (@arrow_rel A1 A2 B1 B2). 176 | 177 | Definition arrow_pointwise_klr A {B1 B2} := 178 | k1 (W:=W) (@arrow_pointwise_rel A B1 B2). 179 | 180 | Definition prod_klr {A1 A2 B1 B2} := 181 | k2 (W:=W) (@prod_rel A1 A2 B1 B2). 182 | 183 | Definition sum_klr {A1 A2 B1 B2} := 184 | k2 (W:=W) (@sum_rel A1 A2 B1 B2). 185 | 186 | Definition list_klr {A1 A2} := 187 | k1 (W:=W) (@list_rel A1 A2). 188 | 189 | Definition set_kle {A B} (R: klr W A B): klr W (A -> Prop) (B -> Prop) := 190 | fun w sA sB => forall a, sA a -> exists b, sB b /\ R w a b. 191 | 192 | Definition set_kge {A B} (R: klr W A B): klr W (A -> Prop) (B -> Prop) := 193 | fun w sA sB => forall b, sB b -> exists a, sA a /\ R w a b. 194 | 195 | Definition klr_union {W A B} := 196 | k2 (W:=W) (@rel_union A B). 197 | 198 | Definition klr_inter {W A B} := 199 | k2 (W:=W) (@rel_inter A B). 200 | 201 | Definition klr_curry {W A1 B1 C1 A2 B2 C2} := 202 | k1 (W:=W) (@rel_curry A1 B1 C1 A2 B2 C2). 203 | End USUAL. 204 | 205 | Notation "RA ==> RB" := (arrow_klr RA RB) 206 | (at level 55, right associativity) : klr_scope. 207 | 208 | Notation "RA ++> RB" := (arrow_klr RA RB) 209 | (at level 55, right associativity) : klr_scope. 210 | 211 | Notation "RA --> RB" := (arrow_klr (k1 flip RA) RB) 212 | (at level 55, right associativity) : klr_scope. 213 | 214 | Notation "- ==> R" := (arrow_pointwise_klr _ R) : klr_scope. 215 | Infix "*" := prod_klr : klr_scope. 216 | Infix "+" := sum_klr : klr_scope. 217 | Infix "\/" := klr_union : klr_scope. 218 | Infix "/\" := klr_inter : klr_scope. 219 | Notation "% R" := (klr_curry R) : klr_scope. 220 | 221 | Global Hint Extern 0 (RIntro _ (arrow_klr _ _ _) _ _) => 222 | eapply k2_rintro : typeclass_instances. 223 | Global Hint Extern 1 (RElim (arrow_klr _ _ _) _ _ _ _) => 224 | eapply k2_relim : typeclass_instances. 225 | 226 | Global Hint Extern 0 (RIntro _ (arrow_pointwise_klr _ _ _) _ _) => 227 | eapply k1_rintro : typeclass_instances. 228 | Global Hint Extern 1 (RElim (arrow_pointwise_klr _ _ _) _ _ _ _) => 229 | eapply k1_relim : typeclass_instances. 230 | 231 | Global Hint Extern 0 (RIntro _ (prod_klr _ _ _) _ _) => 232 | eapply k2_rintro : typeclass_instances. 233 | Global Hint Extern 1 (RElim (prod_klr _ _ _) _ _ _ _) => 234 | eapply k2_relim : typeclass_instances. 235 | 236 | Global Hint Extern 0 (RIntro _ (sum_klr _ _ _) _ _) => 237 | eapply k2_rintro : typeclass_instances. 238 | Global Hint Extern 1 (RElim (sum_klr _ _ _) _ _ _ _) => 239 | eapply k2_relim : typeclass_instances. 240 | 241 | Global Hint Extern 0 (RIntro _ (list_klr _ _) _ _) => 242 | eapply k1_rintro : typeclass_instances. 243 | Global Hint Extern 1 (RElim (list_klr _ _) _ _ _ _) => 244 | eapply k1_relim : typeclass_instances. 245 | 246 | Global Hint Extern 0 (RIntro _ (klr_curry _ _) _ _) => 247 | eapply k1_rintro : typeclass_instances. 248 | Global Hint Extern 1 (RElim (klr_curry _ _) _ _ _ _) => 249 | eapply k1_relim : typeclass_instances. 250 | 251 | (** ** Modal relators *) 252 | 253 | (** In addition to the usual relators defined above, we can define 254 | Kripke-specific relators corresponding to the box and diamond 255 | modalities. *) 256 | 257 | Section MODALITIES. 258 | Context `{kf: KripkeFrame}. 259 | 260 | (** The box modality asserts that the relation holds at all 261 | possible future worlds. *) 262 | 263 | Definition klr_box {A B} (l: L) (R: klr W A B): klr W A B := 264 | fun w x y => forall w', acc l w w' -> R w' x y. 265 | 266 | Global Instance klr_box_subrel {A B}: 267 | Monotonic (@klr_box A B) (- ==> (- ==> subrel) ++> (- ==> subrel)). 268 | Proof. 269 | firstorder. 270 | Qed. 271 | 272 | Lemma klr_box_rintro {A B} l (R: klr W A B) w x y: 273 | RIntro (forall w' (Hw': acc l w w'), R w' x y) (klr_box l R w) x y. 274 | Proof. 275 | firstorder. 276 | Qed. 277 | 278 | Lemma klr_box_relim {A B} l (R: klr W A B) w w' x y P Q: 279 | RElim (R w') x y P Q -> 280 | RElim (klr_box l R w) x y (acc l w w' /\ P) Q. 281 | Proof. 282 | intros H Hxy [Hw' HP]. 283 | apply H; auto. 284 | Qed. 285 | 286 | (** Dually, the diamond modality asserts that the relation holds at 287 | some possible future world. Note the order of the premises in our 288 | intro rule: we want to first determine what [w'] should be, then 289 | attempt to prove [w ~> w']. *) 290 | 291 | Definition klr_diam {A B} (l: L) (R: klr W A B): klr W A B := 292 | fun w x y => exists w', acc l w w' /\ R w' x y. 293 | 294 | Global Instance klr_diam_subrel {A B}: 295 | Monotonic (@klr_diam A B) (- ==> (- ==> subrel) ++> (- ==> subrel)). 296 | Proof. 297 | firstorder. 298 | Qed. 299 | 300 | Lemma klr_diam_rintro {A B} l (R: klr W A B) w w' x y: 301 | RExists (R w' x y /\ acc l w w') (klr_diam l R w) x y. 302 | Proof. 303 | firstorder. 304 | Qed. 305 | 306 | Lemma klr_diam_relim {A B} l (R: klr W A B) w x y P Q: 307 | (forall w', acc l w w' -> RElim (R w') x y P Q) -> 308 | RElim (klr_diam l R w) x y P Q. 309 | Proof. 310 | intros H (w' & Hw' & Hxy) HP. 311 | eapply H; eauto. 312 | Qed. 313 | End MODALITIES. 314 | 315 | Global Instance klr_box_subrel_params: Params (@klr_box) 4 := { }. 316 | Global Instance klr_diam_subrel_params: Params (@klr_diam) 4 := { }. 317 | 318 | Global Hint Extern 0 (RIntro _ (klr_box _ _ _) _ _) => 319 | eapply klr_box_rintro : typeclass_instances. 320 | Global Hint Extern 1 (RElim (klr_box _ _ _) _ _ _ _) => 321 | eapply klr_box_relim : typeclass_instances. 322 | 323 | Global Hint Extern 0 (RExists _ (klr_diam _ _ _) _ _) => 324 | eapply klr_diam_rintro : typeclass_instances. 325 | Global Hint Extern 1 (RElim (klr_diam _ _ _) _ _ _ _) => 326 | eapply klr_diam_relim : typeclass_instances. 327 | 328 | Notation "[ l ] R" := (klr_box l R) (at level 65) : klr_scope. 329 | Notation "< l > R" := (klr_diam l R) (at level 65) : klr_scope. 330 | Notation "[] R" := (klr_box tt R) (at level 65) : klr_scope. 331 | Notation "<> R" := (klr_diam tt R) (at level 65) : klr_scope. 332 | 333 | (** For Kripke frames indexed by pairs, the following variants allow 334 | the components being related to access the labels used for transitions. *) 335 | 336 | Section ARROW_MOD. 337 | Context {W L1 L2} `{kf: KripkeFrame (L1 * L2) W}. 338 | 339 | Definition klr_boxto {A B} (R : klr W A B) : klr W (L1 -> A) (L2 -> B) := 340 | fun w f g => 341 | forall l1 l2 w', acc (l1, l2) w w' -> R w' (f l1) (g l2). 342 | 343 | Definition klr_diamat {A B} (R : klr W A B) : klr W (L1 * A) (L2 * B) := 344 | fun w '(l1, a) '(l2, b) => 345 | exists w', acc (l1, l2) w w' /\ R w' a b. 346 | End ARROW_MOD. 347 | 348 | Notation "[] -> R" := (klr_boxto R) (at level 65) : klr_scope. 349 | Notation "<> * R" := (klr_diamat R) (at level 65) : klr_scope. 350 | 351 | (** ** Flattening KLRs *) 352 | 353 | (** When converting back to a standard [rel], we can quantify over 354 | all worlds. *) 355 | 356 | Section UNKRIPKIFY. 357 | Context `{kf: KripkeFrame}. 358 | 359 | Definition rel_kvd {A B} (R: klr W A B): rel A B := 360 | fun x y => forall w, R w x y. 361 | 362 | Global Instance rel_kvd_subrel A B: 363 | Monotonic (@rel_kvd A B) ((- ==> subrel) ++> subrel). 364 | Proof. 365 | firstorder. 366 | Qed. 367 | 368 | Lemma rel_kvd_rintro {A B} (R: klr W A B) x y: 369 | RIntro (forall w, R w x y) (rel_kvd R) x y. 370 | Proof. 371 | firstorder. 372 | Qed. 373 | 374 | Lemma rel_kvd_relim {A B} (R: klr W A B) w x y P Q: 375 | RElim (R w) x y P Q -> 376 | RElim (rel_kvd R) x y P Q. 377 | Proof. 378 | firstorder. 379 | Qed. 380 | End UNKRIPKIFY. 381 | 382 | Global Instance rel_kvd_subrel_params: Params (@rel_kvd) 3 := { }. 383 | 384 | Global Hint Extern 0 (RIntro _ (rel_kvd _) _ _) => 385 | eapply rel_kvd_rintro : typeclass_instances. 386 | 387 | Global Hint Extern 1 (RElim (rel_kvd _) _ _ _ _) => 388 | eapply rel_kvd_relim : typeclass_instances. 389 | 390 | Notation "|= R" := (rel_kvd R) (at level 65) : rel_scope. 391 | 392 | (** ** Pulling along a Kripke frame morphism *) 393 | 394 | Definition klr_pullw {W1 W2 A B} (f: W1 -> W2) (R: klr W2 A B): klr W1 A B := 395 | fun w => R (f w). 396 | 397 | Notation "R @@ [ f ]" := (klr_pullw f R) 398 | (at level 30, right associativity) : klr_scope. 399 | 400 | Global Instance klr_pullw_subrel {W1 W2 A B} RW1 RW2: 401 | Monotonic 402 | (@klr_pullw W1 W2 A B) 403 | ((RW1 ++> RW2) ++> (RW2 ++> subrel) ++> (RW1 ++> subrel)). 404 | Proof. 405 | unfold klr_pullw. rauto. 406 | Qed. 407 | 408 | Global Instance klr_pullw_subrel_params: 409 | Params (@klr_pullw) 5 := { }. 410 | 411 | Lemma klr_pullw_rintro {W1 W2 A B} (f: W1 -> W2) R w (x:A) (y:B): 412 | RIntro (R (f w) x y) (klr_pullw f R w) x y. 413 | Proof. 414 | firstorder. 415 | Qed. 416 | 417 | Global Hint Extern 0 (RIntro _ (klr_pullw _ _ _) _ _) => 418 | eapply klr_pullw_rintro : typeclass_instances. 419 | 420 | Lemma klr_pullw_relim {W1 W2 A B} (f: W1 -> W2) R w (x:A) (y:B) P Q: 421 | RElim (R (f w)) x y P Q -> 422 | RElim (klr_pullw f R w) x y P Q. 423 | Proof. 424 | firstorder. 425 | Qed. 426 | 427 | Global Hint Extern 1 (RElim (klr_pullw _ _ _) _ _ _ _) => 428 | eapply klr_pullw_relim : typeclass_instances. 429 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | Coqrel was written by Jérémie Koenig, a member of the Yale FLINT Group. 2 | The current version is distributed under the following terms: 3 | 4 | Copyright (c) 2016 Yale University 5 | 6 | Permission is hereby granted, free of charge, to any person obtaining a copy 7 | of this software and associated documentation files (the "Software"), to 8 | deal in the Software without restriction, including without limitation the 9 | rights to use, copy, modify, merge, publish, distribute, sublicense, and/or 10 | sell copies of the Software, and to permit persons to whom the Software is 11 | furnished to do so, subject to the following conditions: 12 | 13 | The above copyright notice and this permission notice shall be included in 14 | all copies or substantial portions of the Software. 15 | 16 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR 17 | IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, 18 | FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE 19 | AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER 20 | LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING 21 | FROM, OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS 22 | IN THE SOFTWARE. 23 | -------------------------------------------------------------------------------- /LogicalRelations.v: -------------------------------------------------------------------------------- 1 | Require Export RelDefinitions. 2 | Require Export RelClasses. 3 | Require Export RelOperators. 4 | Require Export Relators. 5 | Require Export Monotonicity. 6 | Require Export RDestruct. 7 | Require Export MorphismsInterop. 8 | Require Export Transport. 9 | Require Export PreOrderTactic. 10 | -------------------------------------------------------------------------------- /LogicalRelationsTests.v: -------------------------------------------------------------------------------- 1 | Require Import LogicalRelations. 2 | Require Import Coq.Lists.List. 3 | Require Import OptionRel. 4 | Local Open Scope rel_scope. 5 | 6 | (** * Testing Delay.v *) 7 | 8 | Goal 9 | forall P Q : Prop, P -> P /\ Q. 10 | Proof. 11 | Fail delay. 12 | delayed (intros; split; auto; delay); [idtac]. 13 | (* this should leave us with the subgoal Q, which hit [delay] above. *) 14 | Abort. 15 | 16 | (** * Tests *) 17 | 18 | (** ** Reflexivity *) 19 | 20 | Goal 21 | forall A (a: A), a = a. 22 | Proof. 23 | intros. 24 | rauto. 25 | Qed. 26 | 27 | Goal 28 | forall A (a: A), exists b, b = a. 29 | Proof. 30 | intros; eexists. 31 | monotonicity. 32 | Qed. 33 | 34 | (** ** Setoid rewriting *) 35 | 36 | Goal 37 | forall A (a b: A) `(HR: Equivalence A) (H: R a b), 38 | sum_rel R R (inl a) (inl b). 39 | Proof. 40 | intros. 41 | rewrite H. 42 | rewrite <- H. 43 | reflexivity. 44 | Qed. 45 | 46 | (** There is an issue with the following. *) 47 | 48 | Goal 49 | forall A (a b: A) (R: rel A A) (f: A -> A) (p: A -> Prop), 50 | Monotonic f (R ++> R) -> 51 | Monotonic p (R --> impl) -> 52 | R a b -> 53 | p (f b) -> 54 | p (f a). 55 | Proof. 56 | intros A a b R f p Hf Hp Hab H. 57 | rewrite <- Hab in H. 58 | assumption. 59 | Qed. 60 | 61 | (** ** Monotonicity tactics *) 62 | 63 | (** Basic sanity check. This has actually failed in the past due to 64 | [context_candidate] being too liberal and selecting the [RB] 65 | property instead of [RA], then going nowhere with that with no 66 | backtracking implemented yet. *) 67 | 68 | Goal 69 | forall A B (RA: rel A A) (x y: A) (RB: rel B B) (z t: B), 70 | RA x y -> 71 | RB z t -> 72 | RA x y. 73 | Proof. 74 | intros. 75 | rauto. 76 | Qed. 77 | 78 | Goal 79 | forall A (a b: A) (R: rel A A) (H: R a b), 80 | let f (x y: A * A) := (@pair (A+A) (A+A) (inr (fst x)) (inl (snd y))) in 81 | Monotonic f (R * ⊤ ++> ⊤ * R ++> (⊥ + R) * (R + ⊥))%rel. 82 | Proof. 83 | intros; unfold f. 84 | rauto. 85 | Qed. 86 | 87 | Goal 88 | forall {A1 A2 B1 B2} (R1 R1': rel A1 A2) (R2 R2': rel B1 B2), 89 | subrel R1' R1 -> 90 | subrel R2 R2' -> 91 | subrel (R1 ++> R2) (R1' ++> R2'). 92 | Proof. 93 | do 10 intro. 94 | rauto. 95 | Qed. 96 | 97 | (** Check that we can use relational hypotheses from the context as 98 | well as [Monotonic]/[Related] instances. *) 99 | 100 | Goal 101 | forall 102 | {A B} (R: rel A A) 103 | (op: A -> B) (Hop: (R ++> eq) op op) 104 | (x y: A) (Hxy: R x y), 105 | op x = op y. 106 | Proof. 107 | intros. 108 | rauto. 109 | Qed. 110 | 111 | (** Bug with relational parametricity: you can't [RElim] a relation 112 | you don't know yet. *) 113 | 114 | Goal 115 | forall {A B} (RA: rel A A) (RB: rel B B) (m n: (A -> B) * B) (x: A), 116 | ((- ==> RB) * RB)%rel m n -> 117 | RB (fst m x) (fst n x). 118 | Proof. 119 | intros A B RA RB m n x Hmn. 120 | try monotonicity. 121 | try rauto. 122 | Abort. 123 | 124 | (** Pattern matching *) 125 | 126 | Goal 127 | forall {A B} (RA: rel A A) (RB: rel B B) (x y: A) (f: A -> A + B), 128 | RA x y -> 129 | (RA ++> RA + RB) f f -> 130 | RA (match f x with inl a => a | inr b => x end) 131 | (match f y with inl a => a | inr b => y end). 132 | Proof. 133 | intros. 134 | rauto. 135 | Qed. 136 | 137 | Goal 138 | forall {A B} (RA: rel A A) (RB: rel B B) (x y: A * B) (z: A), 139 | RA z z -> 140 | prod_rel RA RB x y -> 141 | RA (let (a, b) := x in z) 142 | (let (a, b) := y in z). 143 | Proof. 144 | intros. 145 | rauto. 146 | Qed. 147 | 148 | Goal 149 | forall {A} (R: rel A A), 150 | Monotonic 151 | (fun (b: bool) x y => if b then x else y) 152 | (- ==> R ++> R ++> R). 153 | Proof. 154 | intros. 155 | rauto. 156 | Qed. 157 | 158 | Goal 159 | forall {A} (R : rel A A) (b : bool) (x y : A), 160 | b = b -> 161 | R x x -> 162 | R y y -> 163 | R (if b then x else y) 164 | (if b then x else y). 165 | Proof. 166 | intros. 167 | rauto. 168 | Qed. 169 | 170 | (** [rel_curry] *) 171 | 172 | Goal 173 | forall {A B C} R R' S (f: A -> B -> B -> C) (x1 y1: A) (x2 y2: B), 174 | Monotonic f (rel_curry (R ++> R' ++> S)) -> 175 | S (f x1 x2 x2) (f y1 y2 y2). 176 | Proof. 177 | intros A B C R R' S f x1 y1 x2 y2 Hf. 178 | monotonicity. 179 | Abort. 180 | 181 | (** If we can deduce some terms are equal, we should be able to 182 | rewrite them under any context. However right now this can only be 183 | done by declaring the [f_equal_relim] hint. 184 | 185 | Maybe a good solution for this might even go further: abstract the 186 | common context and use [eq_rect] if it's possible to prove equality 187 | between the variant parts. *) 188 | 189 | Goal 190 | forall {A} (R: rel A A) (f: A -> A) (C: A -> Prop) x y, 191 | Monotonic f (R ++> eq) -> 192 | R x y -> 193 | C (f x) -> 194 | C (f y). 195 | Proof. 196 | intros A R f C x y Hf Hxy. 197 | Fail rauto. 198 | pose proof @f_equal_relim. 199 | rauto. 200 | Qed. 201 | 202 | (** *** Hypotheses from the context *) 203 | 204 | (* This used to fail because [Hyy] would 205 | shadow [Hxy] (the hypothesis we want). *) 206 | 207 | Goal 208 | forall {A} (R: rel A A) (x y: A), 209 | R x y -> eq y y -> R x y. 210 | Proof. 211 | intros A R x y Hxy Hyy. 212 | monotonicity. 213 | Qed. 214 | 215 | (* This still fail with Coq 8.5, but Coq 8.6 is able to backtrack and 216 | try hypothesis from the context beyond the first one it finds. *) 217 | 218 | Goal 219 | forall {A} (R: rel A A) (x y: A), 220 | R x y -> eq x y -> R x y. 221 | Proof. 222 | intros A R x y Hxy Hyy. 223 | try monotonicity. 224 | Abort. 225 | 226 | (** This used to fail because the flipped hypothesis would not be 227 | identified as a candidate. This is important because the constraints 228 | generated by the setoid rewriting system often have this form. *) 229 | 230 | Goal 231 | forall {A} (R: rel A A) (f : A -> A), 232 | Monotonic f (R ++> R) -> 233 | (flip R ++> flip R) f f. 234 | Proof. 235 | intros A R f Hf. 236 | rauto. 237 | Qed. 238 | 239 | (** Then *this* used to fail because the [flip] was hidden until after 240 | we exploit the [subrel] property. This is important in particular 241 | when using the equivalences for a given partial order, as in [eqrel] 242 | vs. [subrel], where we want to be able to use equivalences in both 243 | directions. *) 244 | 245 | Goal 246 | forall {A} (R R': rel A A) (f: A -> A), 247 | Monotonic f (R ++> R) -> 248 | Related R' (flip R) subrel -> 249 | (R' ++> flip R) f f. 250 | Proof. 251 | intros A R R' f Hf HR. 252 | rauto. 253 | Qed. 254 | 255 | (** As an example where proper handling of [flip] is important, 256 | this is a goal generated by [setoid_rewrite] when we try to rewrite 257 | using [eqrel] under [option_le]. Here, both [R1] and [R2] have to be 258 | instantiated as [flip eq] (so that the associated hypotheses can be 259 | used to solve goals of the form [yi = xi]), and similarly we need to 260 | use the super-relation [flip subrel] of [eqrel]. *) 261 | 262 | Goal 263 | forall A, exists R1 R2: relation (option A), 264 | (eqrel ==> R1 ==> R2 ==> flip impl)%signature option_le option_le. 265 | Proof. 266 | intros A. eexists. eexists. 267 | rauto. 268 | Qed. 269 | 270 | (** *** [impl] vs. [subrel] *) 271 | 272 | (** This checks that a relational property written in terms of 273 | [subrel] can be used to solve a goal stated in terms of [impl]. 274 | This is made possible by [subrel_impl_relim]. *) 275 | 276 | Goal 277 | forall A B C (R: rel A A) (f: A -> rel B C) a1 a2 b c, 278 | Monotonic f (R ++> subrel) -> 279 | R a1 a2 -> 280 | impl (f a1 b c) (f a2 b c). 281 | Proof. 282 | intros A B C R f a1 a2 b c Hf Ha. 283 | monotonicity; rauto. 284 | Qed. 285 | 286 | Goal 287 | forall A1 A2 B1 B2 (R1 R2: rel A1 A2) (R: rel B1 B2), 288 | subrel R1 R2 -> 289 | forall x y, 290 | (R2 ++> R) x y -> 291 | (R1 ++> R) x y. 292 | Proof. 293 | intros A1 A2 B1 B2 R1 R2 R HR12 x y. 294 | rauto. 295 | Qed. 296 | 297 | (** *** Generic rules *) 298 | 299 | (** The [coreflexivity] of [rel_prod] and [eq] makes it possible for 300 | [pair_rel] to behave in the same way as [f_equal] below, since they 301 | allow us to deduce that [eq * eq] is a [subrel] of [eq]. *) 302 | 303 | Goal 304 | forall A B (x1 x2 : A) (y1 y2 : B), 305 | x1 = x2 -> y1 = y2 -> (x1, y1) = (x2, y2). 306 | Proof. 307 | intros. 308 | rauto. 309 | Qed. 310 | 311 | (** ** Using [foo_subrel] instances *) 312 | 313 | (** Still broken because of the interaction between [subrel] and 314 | [- ==> - ==> impl] (or lack thereof) *) 315 | 316 | Goal 317 | forall A1 A2 B1 B2 C1 C2 (R1 R2: rel A1 A2) (R1': rel B1 B2) (R: rel C1 C2), 318 | subrel R1 R2 -> 319 | forall x y, 320 | (R2 ++> R) x y -> 321 | (R1 ++> R) x y. 322 | Proof. 323 | intros A1 A2 B1 B2 C1 C2 R1 R2 R1' R HR12 x y H. 324 | rewrite HR12. 325 | assumption. 326 | Qed. 327 | 328 | Goal 329 | forall A B (xa1 xa2 ya1 ya2 : A) (xb1 xb2 yb1 yb2 : B) 330 | (opA: A -> A -> A) (opB: B -> B -> B) 331 | (RA: rel A A) (RB: rel B B) 332 | (HopA: Monotonic opA (RA ++> RA ++> RA)) 333 | (HopB: Monotonic opB (RB ++> RB ++> RB)) 334 | (Hxa: RA xa1 xa2) 335 | (Hxb: RB xb1 xb2) 336 | (Hya: RA ya1 ya2) 337 | (Hyb: RB yb1 yb2), 338 | (RA * RB)%rel 339 | (opA xa1 ya1, opB xb1 yb1) 340 | (opA xa2 ya2, opB xb2 yb2). 341 | Proof. 342 | intros. 343 | rauto. 344 | Qed. 345 | 346 | Goal 347 | forall A1 A2 B1 B2 C1 C2 (R1 R2: rel A1 A2) (R1': rel B1 B2) (R: rel C1 C2), 348 | subrel R1 R2 -> 349 | forall x y, 350 | (R2 * R1' ++> R) x y -> 351 | (R1 * R1' ++> R) x y. 352 | Proof. 353 | intros A1 A2 B1 B2 C1 C2 R1 R2 R1' R HR12 x y H. 354 | rewrite HR12. 355 | assumption. 356 | Qed. 357 | 358 | (** ** The [rgraph] tactic *) 359 | 360 | Goal 361 | forall {A} (R S T: rel A A), 362 | subrel R S -> 363 | subrel S R -> 364 | subrel S T -> 365 | subrel R T. 366 | Proof. 367 | intros. 368 | rstep. 369 | Qed. 370 | 371 | Goal 372 | forall `(PER) (x y z t : A), 373 | R x y -> 374 | R z y -> 375 | R z t -> 376 | R t x. 377 | Proof. 378 | intros. 379 | rstep. 380 | Qed. 381 | 382 | (** ** General issues *) 383 | 384 | (** Assumptions can be used by [Monotonicity] to solve the goal. 385 | However, consider the goal [?R f g |- (- ==> S) f g]. 386 | [Monotonicity] can solve this by instantiating [R] and using the 387 | hypothesis, however [RIntro] will take precedence in this case. 388 | This means we will end up with a goal [?R f g, x |- S (f x) (g x)] 389 | that cannot be solved anymore. 390 | 391 | When [?R] is already [- ==> S] or some equivalent relation, this is 392 | not a problem, because the work done by [RIntro] can essentially be 393 | undone later by [Monotonicity]. However when [?R] is an existential 394 | variable this is not true and we would prefer an earlier 395 | instantiation. 396 | 397 | This is a possible workaround: 398 | <<< 399 | Lemma assumption_rstep {P : Prop} : 400 | P -> RStep True P. 401 | Proof. 402 | firstorder. 403 | Qed. 404 | 405 | Global Hint Extern 1 (RStep _ (_ ?x ?y)) => 406 | lazymatch goal with 407 | H : ?R x y |- _ => is_evar R; eexact (assumption_rstep H) 408 | end : typeclass_instances. 409 | >>> 410 | *) 411 | 412 | Goal 413 | forall A B (R : relation B) (f g : A -> B), exists X : relation (A -> B), 414 | X f g -> 415 | (- ==> R) f g. 416 | Proof. 417 | intros. eexists. intros. 418 | Fail rauto. 419 | monotonicity. 420 | Abort. 421 | 422 | (** ** The [transport] tactic *) 423 | 424 | Goal 425 | forall W acc A B C (R1: W -> rel A A) (R2: W -> rel B B) (R3: W -> rel C C), 426 | forall f g a b x w, 427 | Monotonic f (rforall w, R1 w ++> R2 w) -> 428 | Monotonic g (rforall w, R2 w ++> option_rel (rel_incr acc R3 w)) -> 429 | R1 w a b -> 430 | g (f a) = Some x -> 431 | exists y, rel_incr acc R3 w x y. 432 | Proof. 433 | intros. 434 | transport H2. 435 | eexists. 436 | rauto. 437 | Qed. 438 | 439 | (** ** Tests for specific relators *) 440 | 441 | (** *** [list_rel] *) 442 | 443 | (** [list_subrel] use to not work because of a missing [Params] declaration. *) 444 | 445 | Goal 446 | forall A B (R R': rel A B) l1 l2 x y, 447 | subrel R R' -> 448 | list_rel R l1 l2 -> 449 | R' x y -> 450 | list_rel R' (x :: l1) (y :: l2). 451 | Proof. 452 | intros. 453 | rauto. 454 | Qed. 455 | 456 | (** *** [rel_pull] *) 457 | 458 | (** The [RIntro] instance for [rel_pull] used to be less general. *) 459 | 460 | Goal 461 | forall A B (f: A -> B) (R: rel B B) x y, 462 | R (f x) (f y) -> 463 | (R @@ f) x y. 464 | Proof. 465 | intros. 466 | rauto. 467 | Qed. 468 | 469 | (** We don't want the introduction rule for [rel_pull] to shadow 470 | relational properties. *) 471 | 472 | Lemma rel_pull_2: 473 | forall A B (f: A -> B) (R: rel B B) (g: A -> A) x y, 474 | Monotonic g (⊤ ==> R @@ f) -> 475 | (R @@ f) (g x) (g y). 476 | Proof. 477 | intros. 478 | rauto. 479 | Qed. 480 | 481 | (** *** [rel_all] *) 482 | 483 | Lemma rel_all_1: 484 | forall {A} (x: A), 485 | (rforall a, req a) x x -> forall a, req a x x. 486 | Proof. 487 | intros. 488 | rauto. 489 | Qed. 490 | 491 | (** *** [rel_ex] *) 492 | 493 | Lemma rel_ex_1: 494 | forall {A} (x: A), 495 | (rexists a, req a) x x. 496 | Proof. 497 | intros. 498 | rauto. 499 | Qed. 500 | 501 | (** *** [rel_top] *) 502 | 503 | (** We used to only have an [RIntro] instance for [⊤%rel], which is 504 | enough for proving all goals of the form [⊤%rel x y] but not enough 505 | to derive properties of relations in which [⊤%rel] is a component. *) 506 | 507 | Lemma rel_top_component_refl: 508 | forall {A} (x: list A), 509 | list_rel ⊤ x x. 510 | Proof. 511 | intros. 512 | rauto. 513 | Qed. 514 | 515 | Lemma rel_top_component_trsym: 516 | forall {A} (x y z: list A), 517 | list_rel ⊤ x y -> list_rel ⊤ z y -> list_rel ⊤ x z. 518 | Proof. 519 | intros. 520 | rauto. 521 | Qed. 522 | -------------------------------------------------------------------------------- /Monotonicity.v: -------------------------------------------------------------------------------- 1 | Require Export RelDefinitions. 2 | Require Export RelOperators. 3 | Require Export Relators. 4 | Require Import Delay. 5 | 6 | (** ** The [monotonicity] tactic *) 7 | 8 | (** The purpose of the [monotonicity] tactic is to automatically 9 | select and apply a theorem of the form [Monotonic ?m ?R] in order to 10 | make progress when the goal is an applied relation. Compared with 11 | setoid rewriting, [monotonicity] is less powerful, but more direct 12 | and simple. This means it is easier to debug, and it can seamlessly 13 | handle dependent types and heterogenous relations. 14 | 15 | The [monotonicity] tactic applies to a goal of the form 16 | [R (f x1 x2 ... xn) (g y1 y2 ... yn)]. The basic idea is to first 17 | locate a relational property declared by the user relating [f] and 18 | [g], namely an instance of [Related f g R'], then to locate and 19 | instance of [RElim] which will allow us to use [R' f g] to make 20 | progress on the current goal. Most of the time, [f] and [g] will be 21 | identical so that the corresponding relational property can be 22 | writted [Monotonic f R']. 23 | 24 | For instance, we can register the monotonicity of [plus] as an 25 | instance of [Monotonic plus (lt ++> lt ++> lt)]. Given an goal of 26 | the form [plus 5 6 < plus x 5], the monotonicity tactic will use 27 | that property and, using the [RElim] rule for [++>], construct an 28 | instance of 29 | [RElim (lt ++> lt ++> lt) plus plus (lt 5 7 -> lt x 5 -> lt (plus 5 6) (plus x 5))] 30 | to apply on the goal, leaving us with the subgoals [5 < 7] and 31 | [x < 5]. 32 | 33 | While this basic principle is rather straightforward, there are some 34 | subtleties that we have to consider in order to build a reliable 35 | system: the relational property we need to use may be about a prefix 36 | [f x1 x2 ... xk] ([k] < [n]) of the application that still has some 37 | arguments applied; we want to be able to exploit [subrel] 38 | relationships to extend the applicability of the declared relational 39 | properties; we need to be need to be careful about how we treat 40 | existential variables. *) 41 | 42 | (** *** Truncating applications *) 43 | 44 | (** The first complication arises from the fact that [f] and [g] may 45 | themselves be partial applications, so the goal may in fact be of 46 | the form 47 | [R (f a1 a2 ... ak x1 x2 ... xn) (g b1 b2 ... bl y1 y2 ... yn)], 48 | with [k] ≠ [l], where the relational property to apply will be of 49 | the form [Related R' (f a1 a2 ... ak) (g b1 b2 ... bl)]. This is 50 | particularly relevant when [f] and [g] may have implicit typeclass 51 | arguments, or type arguments that cannot be covered by our dependent 52 | relators without causing universe inconsistency issues. 53 | 54 | Finding the head terms [f] and [g], then locating a property 55 | [Related R' f g] is rather straightforward, however in this case we 56 | must only drop the last [n] arguments from the applications before 57 | we look up the [Related] instance, with no way of guessing the 58 | correct value of [n]. We want to avoid having to peel off arguments 59 | one by one and performing a [Related] search at each step, so we ask 60 | the user to pitch in and declare instances of [Params f n] and 61 | [Params g n] to guide our search when partial applications with 62 | heads [f] and [g] are involved. 63 | 64 | Note that the term in the goal itself may only be partially 65 | applied. Suppose [f x1 x2 ... xn] is in expecting [p] more 66 | arguments, and that we find an instance of [Params f q]; then we 67 | should only drop [q - p] arguments to obtain the prefix 68 | [f x1 x2 ... x_(n+p-q)]. 69 | 70 | The class [QueryParams m1 m2 n] asserts that the head terms of [m1] 71 | and [m2] have [Params] declarations that instruct us to remove [n] 72 | arguments from each of them before looking up a corresponding 73 | relational property. *) 74 | 75 | Class QueryParams {A B} (m1: A) (m2: B) (n: nat). 76 | 77 | (** If the head term on either side is an existential variable, we 78 | only check [Params] declarations for the other side. This auxiliary 79 | lemma is used to trigger this step; [p] should be the number of 80 | unapplied arguments in whichever term [m1] of [m2] we extracted the 81 | non-existential head term [h] from. *) 82 | 83 | Lemma query_params_one {A B C} (h: A) p (m1: B) (m2: C) n: 84 | Params h (p + n) -> 85 | QueryParams m1 m2 n. 86 | Proof. 87 | constructor. 88 | Qed. 89 | 90 | (** If neither head term is an existential variable, then the [Params] 91 | declarations for both should agree. *) 92 | 93 | Lemma query_params_both {A B C D} (h1: A) (h2: B) p1 p2 (m1: C) (m2: D) n: 94 | Params h1 (p1 + n) -> 95 | Params h2 (p2 + n) -> 96 | QueryParams m1 m2 n. 97 | Proof. 98 | constructor. 99 | Qed. 100 | 101 | (** Now we need to choose which one of these lemmas to apply in a 102 | given situation, and compute the appropriate values for [p]. 103 | We compute [p] by essentially counting the products in the type of 104 | the related terms [m]. Note that we need to make sure we "open up" 105 | types such as [subrel] to expose the products, accomplished here by 106 | using [eval cbv]. *) 107 | 108 | (** FIXME: [count_unapplied_params] fails if [m] has a dependent type, 109 | due to [t'] being a term under binders rather than a closed term. 110 | This means [monotonicity] can't handle goal where the related terms 111 | have dependent types if the corresponding relational property is 112 | about a partial application. *) 113 | 114 | Ltac count_unapplied_params m := 115 | let rec count t := 116 | lazymatch t with 117 | | forall x, ?t' => let n := count t' in constr:(S n) 118 | | _ => constr:(O) 119 | end in 120 | let t := type of m in 121 | let t := eval cbv in t in 122 | count t. 123 | 124 | Ltac query_params m1 m2 := 125 | let rec head t := lazymatch t with ?t' _ => head t' | _ => t end in 126 | let h1 := head m1 in 127 | let p1 := count_unapplied_params m1 in 128 | let h2 := head m2 in 129 | let p2 := count_unapplied_params m2 in 130 | first 131 | [ not_evar h1; not_evar h2; eapply (query_params_both h1 h2 p1 p2) 132 | | not_evar h1; is_evar h2; eapply (query_params_one h1 p1) 133 | | is_evar h1; not_evar h2; eapply (query_params_one h2 p2) ]. 134 | 135 | Global Hint Extern 10 (QueryParams ?m1 ?m2 _) => 136 | query_params m1 m2 : typeclass_instances. 137 | 138 | (** Now that we can figure out how many parameters to drop, we use 139 | that information to construct the pair of prefixes we use to locate 140 | the relational property. We try the following (in this order): 141 | - the whole applications [f x1 ... xn] and [g y1 ... yn]; 142 | - prefixes constructed as above from user-declared [Params] instances; 143 | - the smallest prefix we can obtain by dropping arguments from 144 | both applications simulataneously. 145 | 146 | Again, we need to take into account that either application could 147 | contain an eixstential variable. When we encounter an existential 148 | variable while chopping off arguments, we short-circuit the process 149 | and simply generate a new evar to serve as the shortened version. *) 150 | 151 | Ltac pass_evar_to k := 152 | let Av := fresh "A" in evar (Av : Type); 153 | let A := eval red in Av in clear Av; 154 | let av := fresh "a" in evar (av : A); 155 | let a := eval red in av in clear av; 156 | k a. 157 | 158 | (** The class [RemoveParams] implements the concurrent removal of [n] 159 | parameters from applications [m1] and [m2]. *) 160 | 161 | Class RemoveParams {A B C D} (n: nat) (m1: A) (m2: B) (f: C) (g: D). 162 | 163 | Ltac remove_params_from n m k := 164 | lazymatch n with 165 | | O => k m 166 | | S ?n' => 167 | lazymatch m with 168 | | ?m' _ => remove_params_from n' m' k 169 | | _ => is_evar m; pass_evar_to k 170 | end 171 | end. 172 | 173 | Ltac remove_params n m1 m2 f g := 174 | remove_params_from n m1 ltac:(fun m1' => unify f m1'; 175 | remove_params_from n m2 ltac:(fun m2' => unify g m2')). 176 | 177 | Global Hint Extern 1 (RemoveParams ?n ?m1 ?m2 ?f ?g) => 178 | remove_params n m1 m2 f g; constructor : typeclass_instances. 179 | 180 | (** The class [RemoveAllParams] implements the concurrent removal of 181 | as many arguments as possible. *) 182 | 183 | Class RemoveAllParams {A B C D} (m1: A) (m2: B) (f: C) (g: D). 184 | 185 | Ltac remove_all_params_from m k := 186 | lazymatch m with 187 | | ?m' _ => remove_all_params_from m' k 188 | | _ => not_evar m; k m 189 | end. 190 | 191 | Ltac remove_all_params m1 m2 f g := 192 | first 193 | [ is_evar m2; 194 | remove_all_params_from m1 195 | ltac:(fun m1' => unify f m1'; pass_evar_to ltac:(fun m2' => unify g m2')) 196 | | is_evar m1; 197 | remove_all_params_from m2 198 | ltac:(fun m2' => unify g m2'; pass_evar_to ltac:(fun m1' => unify f m1')) 199 | | lazymatch m1 with ?m1' _ => 200 | lazymatch m2 with ?m2' _ => 201 | remove_all_params m1' m2' f g 202 | end 203 | end 204 | | not_evar m1; unify f m1; 205 | not_evar m2; unify g m2 ]. 206 | 207 | Global Hint Extern 1 (RemoveAllParams ?m1 ?m2 ?f ?g) => 208 | remove_all_params m1 m2 f g; constructor : typeclass_instances. 209 | 210 | (** *** Selecting a relational property *) 211 | 212 | (** With this machinery in place, we can implement the search for 213 | candidate relational properties to use in conjunction with a given 214 | goal [Q]. *) 215 | 216 | Class CandidateProperty {A B} (R': rel A B) f g (Q: Prop) := 217 | candidate_related: R' f g. 218 | 219 | Instance as_is_candidate {A B} (R R': rel A B) m1 m2: 220 | Related m1 m2 R' -> 221 | CandidateProperty R' m1 m2 (R m1 m2) | 10. 222 | Proof. 223 | firstorder. 224 | Qed. 225 | 226 | Instance remove_params_candidate {A B C D} R (m1: A) (m2: B) R' (f: C) (g: D) n: 227 | QueryParams m1 m2 n -> 228 | RemoveParams n m1 m2 f g -> 229 | Related f g R' -> 230 | CandidateProperty R' f g (R m1 m2) | 20. 231 | Proof. 232 | firstorder. 233 | Qed. 234 | 235 | Instance remove_all_params_candidate {A B C D} R (m1:A) (m2:B) R' (f:C) (g:D): 236 | RemoveAllParams m1 m2 f g -> 237 | Related f g R' -> 238 | CandidateProperty R' f g (R m1 m2) | 30. 239 | Proof. 240 | firstorder. 241 | Qed. 242 | 243 | (** We also attempt to use any matching hypothesis. This takes 244 | priority and bypasses any [Params] declaration. We assume that we 245 | will always want such hypotheses to be used, at least when the left- 246 | or right-hand side matches exactly. There is a possibility that this 247 | ends up being too broad for some applications, for which we'll want 248 | to restrict ourselves to [Related] instances explicitely defined by 249 | the user. If this turns out to be the case, we can introduce an 250 | intermediate class with more parameters to control the sources of 251 | the relational properties we use, and perhaps have some kind of 252 | normalization process akin to what is done in [Coq.Classes.Morphisms]. 253 | 254 | We don't hesitate to instantiate evars in the goal or hypothesis; if 255 | the types or skeletons are compatible this is likely what we want. 256 | In one practical case where this is needed, my hypothesis is 257 | something like [R (f ?x1 ?y1 42) (f ?x2 ?y2 24)], and I want 258 | [?x1], [?x2], [?y1], [?y2] to be matched against the goal. In case 259 | the use of [unify] below is too broad, we could develop a strategy 260 | whereby at least one of the head terms [f], [g] should match exactly. 261 | 262 | We also consider the possibility that the relation's arguments may 263 | be flipped in the hypothesis. In that case we wrap the relation with 264 | [flip] to construct a corresponding [CandidateProperty]. This [flip] 265 | will most likely be handled at the [RElim] stage by [flip_relim]. 266 | The hypothesis relation itself may be of the form [flip R]. In that 267 | case, the [CandidateProperty] will use [flip (flip R)] and 268 | [flip_relim] will simply be used twice. In more common and subtle 269 | variation, the hypothesis relation is not directly of this form, but 270 | should either be instantiated or refined (through [subrel]) into 271 | such a form. This is handled by [rimpl_flip_subrel]. 272 | 273 | Such issues are especially common when solving a goal of the 274 | form [(R --> R') f g], or goals generated by the setoid rewriting 275 | system of the form [Proper (?R1 ==> ... ==> ?Rn ==> flip impl)], 276 | and where the [?Ri]s will generally need to be instantiated as 277 | [flip ?Ri']. We do not pursue a similar strategy when looking up 278 | candidates from the typeclass instance database, because we expect 279 | the user to normalize such properties before they declare them, for 280 | example declare an instance of [Related (R1 ++> R2 --> R3) g f] 281 | rather than an instance of [Related (R1 --> R2 ++> flip R3) f g]. 282 | 283 | Note that it is important that we reduce the goal to [?R ?m ?n] 284 | before we use [eexact]: if the relation in the hypothesis is an 285 | existential variable, we don't want it unified against 286 | [CandidateProperty _ _ _ (_ ?m ?n)]. *) 287 | 288 | Lemma flip_context_candidate {A B} (R: rel A B) x y : R y x -> flip R x y. 289 | Proof. 290 | firstorder. 291 | Qed. 292 | 293 | Ltac context_candidate := 294 | let rec is_prefix f m := 295 | first 296 | [ unify f m 297 | | lazymatch m with ?n _ => is_prefix f n end ] in 298 | let rec is_prefixable f m := 299 | first 300 | [ is_evar f 301 | | is_evar m 302 | | unify f m 303 | | lazymatch m with ?n _ => is_prefixable f n end ] in 304 | multimatch goal with 305 | | H: _ ?f ?g |- @CandidateProperty ?A ?B ?R ?x ?y (_ ?m ?n) => 306 | red; 307 | once first 308 | [ is_prefix f m; is_prefixable g n 309 | | is_prefix g n; is_prefixable f m 310 | | is_prefix g m; is_prefixable f n; apply flip_context_candidate 311 | | is_prefix f n; is_prefixable g m; apply flip_context_candidate ]; 312 | eexact H 313 | end. 314 | 315 | Global Hint Extern 1 (CandidateProperty _ _ _ _) => 316 | context_candidate : typeclass_instances. 317 | 318 | (** *** Using [subrel] *) 319 | 320 | (** It is not obvious at what point in the process [subrel] should be 321 | hooked in. One thing we crucially want to avoid is an open-ended 322 | [subrel] search enumerating all possibilities to be filtered later, 323 | with a high potential for exponential blow-up should the user be a 324 | little too liberal with the [subrel] instances they declare. 325 | 326 | Here I choose to have it kick in after a candidate property has been 327 | selected, and we know how to apply it to a goal. Then we use 328 | [subrel] to bridge any gap between that goal and the actual one, 329 | through the [RImpl] class below. 330 | 331 | This is a conservative solution which precludes many interesting 332 | usages of [subrel]. For instance, suppose we have a relational 333 | property alogn the lines of [Monotonic f ((R1 ++> R1) ∩ (R2 ++> R2))]. 334 | We would want to be able to use it to show that [f] preserve [R1] or 335 | [R2] individually (because [subrel (R1 ++> R1) ((R1 ++> R1) ∩ (R2 336 | ++> R2))], but also together (because [subrel (R1 ∩ R2 ++> R1 ∩ R2) 337 | ((R1 ++> R1) ∩ (R2 ++> R2))]). This cannot be done using this 338 | approach, which does not act on the relational property itself but 339 | only the goal we're attempting to prove. 340 | 341 | Perhaps in the future we can extend this by acting at the level of 342 | [RElim]. In any case, we should provide explicit guidelines for when 343 | to declare [subrel] instances, and how. *) 344 | 345 | Class RImpl (P Q: Prop): Prop := 346 | rimpl: P -> Q. 347 | 348 | (** While we give priority to [rimpl_refl] below, when it doesn't work 349 | we use [RAuto] to establish a [subrel] property. The [Monotonic] 350 | instances we declared alongside relators can be used conjunction 351 | with [Monotonicity] to break up [subrel] goals along the structure 352 | of the relations being considered. This may end up causing loops 353 | (especially in failure cases), so it may be necessary to add some 354 | flag in the context to prevent [rimpl_subrel] from being used when 355 | discharging the [subrel] goals themselves. *) 356 | 357 | Global Instance rimpl_refl {A B} (R : rel A B) m n: 358 | RImpl (R m n) (R m n). 359 | Proof. 360 | firstorder. 361 | Qed. 362 | 363 | Global Instance rimpl_subrel {A B} (R R': rel A B) m n: 364 | NonDelayed (RAuto (subrel R R')) -> 365 | RImpl (R m n) (R' m n). 366 | Proof. 367 | firstorder. 368 | Qed. 369 | 370 | Global Instance rimpl_flip_subrel {A B} R R' (x: A) (y: B): 371 | NonDelayed (RAuto (subrel R (flip R'))) -> 372 | RImpl (R x y) (R' y x) | 2. 373 | Proof. 374 | firstorder. 375 | Qed. 376 | 377 | (** *** Main tactic *) 378 | 379 | (** With these components, defining the [monotonicity] tactic is 380 | straightforward: identify a candidate property, then figure out a 381 | way to apply it to the goal [Q] using the [RElim] class. We first 382 | define a [Monotonicity] typeclass that captures this behavior with 383 | full backtracking ability. *) 384 | 385 | Class Monotonicity (P Q: Prop): Prop := 386 | monotonicity: P -> Q. 387 | 388 | Global Instance apply_candidate {A B} (R: rel A B) m n P Q Q': 389 | CandidateProperty R m n Q -> 390 | RElim R m n P Q' -> 391 | RImpl Q' Q -> 392 | Monotonicity P Q. 393 | Proof. 394 | firstorder. 395 | Qed. 396 | 397 | (** The Ltac tactic simply applies [monotonicity]; typeclass 398 | resolution will do the rest. Note that using [apply] naively is too 399 | lenient because in a goal of type [A -> B], it will end up unifying 400 | [A] with [P] and [B] with [Q] instead of setting [Q := A -> B] and 401 | generating a new subgoal for [P] as expected. On the other hand, 402 | using [refine] directly is too restrictive because it will not unify 403 | the type of [monotonicity] against the goal if existential variables 404 | are present in one or the other. Hence we constrain apply just 405 | enough, so as to handle both of these cases. *) 406 | 407 | Ltac monotonicity := 408 | lazymatch goal with |- ?Q => apply (monotonicity (Q:=Q)) end; 409 | Delay.split_conjunction. 410 | 411 | (** Another way to use [Monotonicity] is to hook it as an [RStep] 412 | instance. *) 413 | 414 | Global Instance monotonicity_rstep {A B} (P: Prop) (R: rel A B) m n: 415 | PolarityResolved R -> 416 | Monotonicity P (R m n) -> 417 | RStep P (R m n) | 50. 418 | Proof. 419 | firstorder. 420 | Qed. 421 | 422 | (** *** [subrel] constraint as a subgoal *) 423 | 424 | (** When [monotonicity] fails it can be hard to figure out what is 425 | going wrong, especially if the problem occurs in the [subrel] goal 426 | of [rimpl_subrel]. The following tactic emits this constraint as a 427 | subgoal instead of solving it with [RAuto], so that the user can 428 | figure out what's broken and/or prove it manually in contexts where 429 | an automatic resolution is tricky. *) 430 | 431 | Class SubrelMonotonicity (P Q: Prop): Prop := 432 | subrel_monotonicity: P -> Q. 433 | 434 | Global Instance apply_candidate_subrel {A B C D} (R: rel A B) m n P Rc Rg x y: 435 | CandidateProperty R m n (Rg x y) -> 436 | RElim R m n P (Rc x y) -> 437 | SubrelMonotonicity (@subrel C D Rc Rg /\ P) (Rg x y). 438 | Proof. 439 | firstorder. 440 | Qed. 441 | 442 | Ltac subrel_monotonicity := 443 | lazymatch goal with |- ?Q => apply (subrel_monotonicity (Q:=Q)) end; 444 | Delay.split_conjunction. 445 | 446 | 447 | (** ** Generic instances *) 448 | 449 | (** When all else fail, we would like to fall back on a behavior 450 | similar to that of [f_equal]. To that end, we add a default 451 | [CandidateProperty] that identifies the longest common prefix of the 452 | applications in the goal and asserts that they are equal. *) 453 | 454 | Class CommonPrefix {A B C} (m1: A) (m2: B) (f: C). 455 | 456 | Ltac common_prefix m1 m2 f := 457 | first 458 | [ not_evar m1; not_evar m2; 459 | unify m1 m2; unify f m1 460 | | lazymatch m1 with ?m1' _ => 461 | lazymatch m2 with ?m2' _ => 462 | common_prefix m1' m2' f 463 | end 464 | end 465 | | unify m1 m2; unify f m1 ]. 466 | 467 | Global Hint Extern 1 (CommonPrefix ?m1 ?m2 ?f) => 468 | common_prefix m1 m2 f; constructor : typeclass_instances. 469 | 470 | Global Instance eq_candidate {A B C} R (m1: A) (m2: B) (f: C): 471 | CommonPrefix m1 m2 f -> 472 | CandidateProperty eq f f (R m1 m2) | 100. 473 | Proof. 474 | firstorder. 475 | Qed. 476 | 477 | (** Then, we can use the following [RElim] instance to obtain the 478 | behavior of the [f_equal] tactic. 479 | 480 | In practice, I find that [f_equal_relim] is too broadly applicable. 481 | In situations where a relational property is missing or inapplicable 482 | for some reason, we would like [repeat rstep] to leave us in a 483 | situation where we can examine what is going wrong, or do a manual 484 | proof when necessary. If we use [f_equal]-like behavior as a 485 | default, we will instead be left several step further, where it is 486 | no longer obvious which function is involved, and where the goal may 487 | not be provable anymore. 488 | 489 | With that said, if you would like to switch [f_equal] on, you can 490 | register [f_equal_elim] as an instance with the following hint: 491 | << 492 | Hint Extern 1 (RElim (@eq (_ -> _)) _ _ _ _) => 493 | eapply f_equal_relim : typeclass_instances. 494 | >> *) 495 | 496 | Lemma f_equal_relim {A B} f g m n P Q: 497 | RElim eq (f m) (g n) P Q -> 498 | RElim (@eq (A -> B)) f g (m = n /\ P) Q. 499 | Proof. 500 | intros Helim Hf [Hmn HP]. 501 | apply Helim; eauto. 502 | congruence. 503 | Qed. 504 | 505 | (** Note that thanks to [eq_subrel], this will also apply to a goal 506 | that uses a [Reflexive] relation, not just a goal that uses [eq]. 507 | In fact, this subsumes the [reflexivity] tactic as well, which 508 | corresponds to the special case where all arguments are equal; 509 | in that case [relim_base] can be used directly and [f_equal_elim] is 510 | not necessary. *) 511 | -------------------------------------------------------------------------------- /MorphismsInterop.v: -------------------------------------------------------------------------------- 1 | Require Import Coq.Classes.Morphisms. 2 | Require Import Coq.Relations.Relation_Definitions. 3 | Require Import RelDefinitions. 4 | Require Import RelOperators. 5 | Require Import Relators. 6 | Require Import Monotonicity. 7 | 8 | (** * Interoperation with [Coq.Classes.Morphisms] *) 9 | 10 | (** Although ultimately we will reimplement a [rewrite] tactic, and 11 | our system is largely designed to replace the setoid rewriting 12 | support in [Coq.Classes.Morphisms], we still want to be compatible 13 | with it so that: 14 | 15 | - we can take advantage of the relational properties declared in 16 | the old system; 17 | - the generalized [setoid_rewrite] can take advantage of our 18 | library for answering its [Morphisms.Proper] queries. 19 | 20 | This compatibility layer is pretty straightforward: we simply 21 | declare the instances we need in order to use the properties 22 | declared in the old system, and use [rauto] to handle 23 | [Morphisms.Proper] queries. *) 24 | 25 | (** ** Using [Morphisms]-style relational properties *) 26 | 27 | (** *** Relators *) 28 | 29 | (** While [Coq.Classes.Morphisms] uses a distinct set of relators from 30 | ours, we don't try to map them into our language. Instead, we hook 31 | them into our system the same way we do our generalized versions, 32 | so that [respectful] and the like can be used as just another set of 33 | "native" Coqrel relators. *) 34 | 35 | Global Instance respectful_subrel A B: 36 | Monotonic (@respectful A B) (subrel --> subrel ++> subrel). 37 | Proof. 38 | firstorder. 39 | Qed. 40 | 41 | Global Instance respectful_params: 42 | Params (@respectful) 4 := { }. 43 | 44 | Lemma respectful_rintro {A B} (RA: relation A) (RB: relation B) f g: 45 | RIntro (forall x y, RA x y -> RB (f x) (g y)) (respectful RA RB) f g. 46 | Proof. 47 | firstorder. 48 | Qed. 49 | 50 | Global Hint Extern 0 (RIntro _ (respectful _ _) _ _) => 51 | eapply respectful_rintro : typeclass_instances. 52 | 53 | Lemma respectful_relim {A B} (RA: relation A) (RB: relation B) f g m n P Q: 54 | RElim RB (f m) (g n) P Q -> 55 | RElim (respectful RA RB) f g (RA m n /\ P) Q. 56 | Proof. 57 | firstorder. 58 | Qed. 59 | 60 | Global Hint Extern 1 (RElim (respectful _ _) _ _ _ _) => 61 | eapply respectful_relim : typeclass_instances. 62 | 63 | Global Instance forall_relation_subrel A P: 64 | Monotonic (@forall_relation A P) ((forallr -, subrel) ++> subrel). 65 | Proof. 66 | firstorder. 67 | Qed. 68 | 69 | Global Instance forall_relation_params: 70 | Params (@forall_relation) 3 := { }. 71 | 72 | Lemma forall_relation_rintro {A B} (R: forall a:A, relation (B a)) f g: 73 | RIntro (forall a, R a (f a) (g a)) (forall_relation R) f g. 74 | Proof. 75 | firstorder. 76 | Qed. 77 | 78 | Global Hint Extern 0 (RIntro _ (forall_relation _) _ _) => 79 | eapply forall_relation_rintro : typeclass_instances. 80 | 81 | Lemma forall_relation_relim {A B} (R: forall a:A, relation (B a)) f g a P Q: 82 | RElim (R a) (f a) (g a) P Q -> 83 | RElim (forall_relation R) f g P Q. 84 | Proof. 85 | firstorder. 86 | Qed. 87 | 88 | Global Hint Extern 1 (RElim (forall_relation _) _ _ _ _) => 89 | eapply forall_relation_relim : typeclass_instances. 90 | 91 | Global Instance pointwise_relation_subrel A B: 92 | Monotonic (@pointwise_relation A B) (subrel ++> subrel). 93 | Proof. 94 | firstorder. 95 | Qed. 96 | 97 | Global Instance pointwise_relation_params: 98 | Params (@pointwise_relation) 3 := { }. 99 | 100 | Lemma pointwise_relation_rintro {A B} (R: relation B) f g: 101 | RIntro (forall a, R (f a) (g a)) (pointwise_relation A R) f g. 102 | Proof. 103 | firstorder. 104 | Qed. 105 | 106 | Global Hint Extern 0 (RIntro _ (pointwise_relation _ _) _ _) => 107 | eapply pointwise_relation_rintro : typeclass_instances. 108 | 109 | Lemma pointwise_relation_relim {A B} (R: relation B) f g a P Q: 110 | RElim R (f a) (g a) P Q -> 111 | RElim (pointwise_relation A R) f g P Q. 112 | Proof. 113 | firstorder. 114 | Qed. 115 | 116 | Global Hint Extern 1 (RElim (pointwise_relation _ _) _ _ _ _) => 117 | eapply pointwise_relation_relim : typeclass_instances. 118 | 119 | (** The rewriting system sometimes generates arguments constraints of 120 | the following form (when rewriting an argument of [arrow_rel] for 121 | instance), which need to be matched against our relational 122 | properties stated in terms of [subrel]. *) 123 | 124 | Global Instance pointwise_relation_subrel_subrel A B: 125 | Related (pointwise_relation A (pointwise_relation B impl)) subrel subrel. 126 | Proof. 127 | firstorder. 128 | Qed. 129 | 130 | (** *** [Morphisms.Proper] instances *) 131 | 132 | (** Now that we can interpret the relators used in [Morphisms]-style 133 | relational properties, we can simply hook instances of 134 | [Morphisms.Proper] into our own database. 135 | 136 | We have to be careful about how this instance works. For one thing, 137 | we are not interested in going through the whole resolution process 138 | in [Coq.Classes.Morphisms], but are only interested in recovering 139 | user-defined instances as-is. Moreover, we want to avoid a loop with 140 | [solve_morphisms_proper] below, where our library (and hence 141 | [Related] instances) are used to solve [Proper] queries. 142 | We can make sure both of those things are avoided by using a 143 | [normalization_done] hypothesis with our [Proper] query. *) 144 | 145 | Global Instance morphisms_proper_related {A} (R: relation A) (a: A): 146 | (normalization_done -> Proper R a) -> 147 | Monotonic a R | 10. 148 | Proof. 149 | firstorder. 150 | Qed. 151 | 152 | (** *** [subrelation] instances *) 153 | 154 | (** Likewise, we want to be able to use [subrelation] declaration as 155 | [subrel] instances. Again, we use an immediate hint. *) 156 | 157 | Lemma subrelation_subrel {A} (R1 R2: relation A): 158 | Unconvertible R1 R2 -> 159 | subrelation R1 R2 -> 160 | Related R1 R2 subrel. 161 | Proof. 162 | firstorder. 163 | Qed. 164 | 165 | Global Hint Immediate subrelation_subrel : typeclass_instances. 166 | 167 | 168 | (** ** Satisfying [Morphisms.Proper] queries *) 169 | 170 | Ltac solve_morphisms_proper := 171 | match goal with 172 | | _ : normalization_done |- _ => 173 | fail 1 174 | | H : apply_subrelation |- _ => 175 | clear H; 176 | red; 177 | rauto 178 | end. 179 | 180 | Global Hint Extern 10 (Morphisms.Proper _ _) => 181 | solve_morphisms_proper : typeclass_instances. 182 | -------------------------------------------------------------------------------- /OptionRel.v: -------------------------------------------------------------------------------- 1 | Require Import LogicalRelations. 2 | 3 | 4 | (** * The [option_le] relator *) 5 | 6 | Inductive option_le {A1 A2} (RA: rel A1 A2): rel (option A1) (option A2) := 7 | | Some_le_def x y: RA x y -> option_le RA (Some x) (Some y) 8 | | None_le_def y: option_le RA None y. 9 | 10 | Global Instance Some_le: 11 | Monotonic (@Some) (forallr R @ A1 A2 : rel, R ++> option_le R). 12 | Proof. 13 | exact @Some_le_def. 14 | Qed. 15 | 16 | Lemma None_le {A B} R y: 17 | RIntro True (@option_le A B R) None y. 18 | Proof. 19 | intros _. 20 | apply @None_le_def. 21 | Qed. 22 | 23 | Global Hint Extern 0 (RIntro _ (option_le _) None _) => 24 | eapply None_le : typeclass_instances. 25 | 26 | Global Instance option_le_subrel A B: 27 | Monotonic (@option_le A B) (subrel ++> subrel). 28 | Proof. 29 | intros R1 R2 HR x y Hxy. 30 | destruct Hxy; constructor; eauto. 31 | Qed. 32 | 33 | Global Instance option_le_subrel_params: 34 | Params (@option_le) 3 := { }. 35 | 36 | Global Instance option_le_rel {A B}: 37 | Related (@option_rel A B) (@option_le A B) (subrel ++> subrel) | 10. 38 | Proof. 39 | intros R1 R2 HR _ _ []; constructor; eauto. 40 | Qed. 41 | 42 | Lemma option_le_refl {A} (R: relation A): 43 | Reflexive R -> Reflexive (option_le R). 44 | Proof. 45 | intros H. 46 | intros [x|]; constructor; reflexivity. 47 | Qed. 48 | 49 | Global Hint Extern 1 (Reflexive (option_le ?R)) => 50 | eapply option_le_refl : typeclass_instances. 51 | 52 | Lemma option_le_trans {A} (R: relation A): 53 | Transitive R -> Transitive (option_le R). 54 | Proof. 55 | intros H. 56 | intros _ _ z [x y Hxy | y] Hz; inversion Hz; subst; clear Hz. 57 | - constructor. 58 | etransitivity; eassumption. 59 | - constructor. 60 | - constructor. 61 | Qed. 62 | 63 | Global Hint Extern 1 (Transitive (option_le ?R)) => 64 | eapply option_le_trans : typeclass_instances. 65 | 66 | Global Instance option_map_le: 67 | Monotonic 68 | (@option_map) 69 | (forallr RA, forallr RB, (RA ++> RB) ++> option_le RA ++> option_le RB). 70 | Proof. 71 | intros A1 A2 RA B1 B2 RB f g Hfg x y Hxy. 72 | destruct Hxy; constructor; eauto. 73 | Qed. 74 | 75 | 76 | (** * The [option_ge] relator *) 77 | 78 | Inductive option_ge {A1 A2} (RA: rel A1 A2): rel (option A1) (option A2) := 79 | | Some_ge_def x y: RA x y -> option_ge RA (Some x) (Some y) 80 | | None_ge_def x: option_ge RA x None. 81 | 82 | Global Instance Some_ge: 83 | Monotonic (@Some) (forallr R @ A1 A2 : rel, R ++> option_ge R). 84 | Proof. 85 | exact @Some_ge_def. 86 | Qed. 87 | 88 | Lemma None_ge {A B} R x: 89 | RIntro True (@option_ge A B R) x None. 90 | Proof. 91 | intros _. 92 | apply @None_ge_def. 93 | Qed. 94 | 95 | Global Hint Extern 0 (RIntro _ (option_ge _) _ None) => 96 | eapply None_ge : typeclass_instances. 97 | 98 | Global Instance option_ge_subrel A B: 99 | Monotonic (@option_ge A B) (subrel ++> subrel). 100 | Proof. 101 | intros R1 R2 HR x y Hxy. 102 | destruct Hxy; constructor; eauto. 103 | Qed. 104 | 105 | Global Instance option_ge_subrel_params: 106 | Params (@option_ge) 3 := { }. 107 | 108 | Global Instance option_ge_rel {A B}: 109 | Related (@option_rel A B) (@option_ge A B) (subrel ++> subrel) | 10. 110 | Proof. 111 | intros R1 R2 HR _ _ []; constructor; eauto. 112 | Qed. 113 | 114 | Lemma option_ge_refl {A} (R: relation A): 115 | Reflexive R -> Reflexive (option_ge R). 116 | Proof. 117 | intros H. 118 | intros [x|]; constructor; reflexivity. 119 | Qed. 120 | 121 | Global Hint Extern 1 (Reflexive (option_ge ?R)) => 122 | eapply option_ge_refl : typeclass_instances. 123 | 124 | Lemma option_ge_trans {A} (R: relation A): 125 | Transitive R -> Transitive (option_ge R). 126 | Proof. 127 | intros H. 128 | intros _ _ z [x y Hxy | y] Hz; inversion Hz; subst; clear Hz. 129 | - constructor. 130 | etransitivity; eassumption. 131 | - constructor. 132 | - constructor. 133 | Qed. 134 | 135 | Global Hint Extern 1 (Transitive (option_ge ?R)) => 136 | eapply option_ge_trans : typeclass_instances. 137 | 138 | Global Instance option_map_ge: 139 | Monotonic 140 | (@option_map) 141 | (forallr RA, forallr RB, (RA ++> RB) ++> option_ge RA ++> option_ge RB). 142 | Proof. 143 | intros A1 A2 RA B1 B2 RB f g Hfg x y Hxy. 144 | destruct Hxy; constructor; eauto. 145 | Qed. 146 | 147 | 148 | (** * [Transport] instances for [option] relations *) 149 | 150 | Global Instance option_le_transport_eq_some {A B} (R: rel A B) x y a: 151 | Transport (option_le R) x y (x = Some a) (exists b, y = Some b /\ R a b). 152 | Proof. 153 | intros Hxy Hx. 154 | subst; inversion Hxy; eauto. 155 | Qed. 156 | 157 | Global Instance option_ge_transport_eq_none {A B} (R: rel A B) x y: 158 | Transport (option_ge R) x y (x = None) (y = None). 159 | Proof. 160 | intros Hxy Hx. 161 | subst; inversion Hxy; eauto. 162 | Qed. 163 | -------------------------------------------------------------------------------- /PreOrderTactic.v: -------------------------------------------------------------------------------- 1 | Require Import RelDefinitions. 2 | 3 | (** For a [Transitive] and possibly [Symmetric] relation [R], the 4 | following tactic attempts to use hypotheses from the context to 5 | prove a goal of the form [R m n] through a simple depth-first graph 6 | search algorithm as follows. 7 | 8 | We use [m] as the starting node. We maintain a stack of terms [t] 9 | together with proofs that for any [x], if [R t x] then also [R m x]. 10 | The stack is initialized to [(m, fun x H => H) :: nil]. Then as long 11 | as the stack is non-empty, pop the first element [(t, Ht)], and for 12 | every hypothesis of the form [Htt' : R t ?t']: 13 | 14 | - If [t'] is [n], success! Use [Ht t' Htt']. 15 | - Otherwise, push [(t', fun x Hx => transitivity (Ht t' Htt') Hx)] 16 | on the stack. 17 | 18 | If the relation is symmetric, repeat with hypotheses of the form 19 | [R ?t' t], then go on to pop the next element. If the search 20 | terminates without encountering [n], fail. Once a hypothesis has 21 | been used, it is "deactivated" so as not to be used again, using the 22 | following marker. *) 23 | 24 | Definition rgraph_done {A B} (R: rel A B) a b := R a b. 25 | 26 | (** In addition, we roll our own data structure to make sure we avoid 27 | any problems related to universe inconsistencies. *) 28 | 29 | Inductive rgraph_queue {A} (R: rel A A) (m: A) := 30 | | rgraph_queue_nil: 31 | rgraph_queue R m 32 | | rgraph_queue_cons x: 33 | (forall y, R x y -> R m y) -> 34 | rgraph_queue R m -> 35 | rgraph_queue R m. 36 | 37 | (** Notes: 1. it's not clear that the [not_evar] checks below are 38 | actually necessary. Nothing here should instantiate any evar. They 39 | can probably be treated like any other term. 2. Ideally, we 40 | probably want to do the graph traversal first, then look for 41 | instances of [Transitive] and [Symmetric], since the number of steps 42 | in the traversal is bounded by the number of hypotheses, whereas a 43 | typeclass resolution query is a rather open-ended thing. 3. With 44 | this in mind, we can likely first do a traversal without symmetry 45 | and see where we get, then if necessary switch symmetry on at that 46 | point and try to reach more vertices this way. *) 47 | 48 | Ltac rgraph := 49 | lazymatch goal with 50 | | |- ?R ?m ?n => 51 | not_evar R; 52 | not_evar m; 53 | not_evar n; 54 | first 55 | [ assert (Transitive R) by typeclasses eauto 56 | | fail 1 R "is not declared transitive" ]; 57 | try assert (Symmetric R) by typeclasses eauto; 58 | let rec step q := 59 | lazymatch q with 60 | | rgraph_queue_cons _ _ ?t ?Ht ?tail => 61 | gather tail t Ht 62 | end 63 | with gather q t Ht := 64 | let push u Htu := 65 | let Hu := constr:(fun x Hux => transitivity (z:=x) (Ht u Htu) Hux) in 66 | gather (rgraph_queue_cons R m u Hu q) t Ht in 67 | lazymatch goal with 68 | | Htn : R t n |- _ => 69 | exact (Ht n Htn) 70 | | Hnt : R n t, HRsym: Symmetric R |- _ => 71 | exact (Ht n (symmetry (Symmetric := HRsym) Hnt)) 72 | | Htu: R t ?u |- _ => 73 | change (rgraph_done R t u) in Htu; 74 | push u Htu 75 | | Hut: R ?u t, HRsym: Symmetric R |- _ => 76 | change (rgraph_done R u t) in Hut; 77 | push u (symmetry (Symmetric := HRsym) Hut) 78 | | _ => 79 | step q 80 | end in 81 | first 82 | [ gather (rgraph_queue_nil R m) m (fun x (H: R m x) => H) 83 | | fail 1 n "not reachable from" m "using hypotheses from the context" ] 84 | | _ => 85 | fail "the goal is not an applied relation" 86 | end. 87 | 88 | (** Hook [rgraph] for providing [RStep]s *) 89 | 90 | Global Hint Extern 30 (RStep ?P (_ _ _)) => 91 | solve [ try unify P True; intro; rgraph ] : typeclass_instances. 92 | -------------------------------------------------------------------------------- /RDestruct.v: -------------------------------------------------------------------------------- 1 | Require Export RelDefinitions. 2 | 3 | (** * The [RDestruct] class *) 4 | 5 | (** ** Tactics *) 6 | 7 | (** The [rdestruct] tactic locates and uses an [RDestruct] instance 8 | for a given hypothesis. *) 9 | 10 | Ltac rdestruct H := 11 | lazymatch type of H with 12 | | ?R ?m ?n => 13 | not_evar R; 14 | pattern m, n; 15 | apply (rdestruct (R:=R) m n H); 16 | clear H; 17 | Delay.split_conjunction 18 | end. 19 | 20 | (** The [rinversion] tactic is similar, but it remembers the terms 21 | related by the destructed hypothesis. *) 22 | 23 | Ltac rinversion_tac H Hm Hn := 24 | lazymatch type of H with 25 | | ?R ?m ?n => 26 | not_evar R; 27 | pattern m, n; 28 | lazymatch goal with 29 | | |- ?Q _ _ => 30 | generalize (eq_refl m), (eq_refl n); 31 | change ((fun x y => Delay.delayed_goal (x = m -> y = n -> Q x y)) m n); 32 | apply (rdestruct (R:=R) m n H); 33 | Delay.split_conjunction; 34 | intros Hm Hn 35 | end 36 | end. 37 | 38 | Tactic Notation "rinversion" constr(H) "as" ident(Hl) "," ident(Hr) := 39 | rinversion_tac H Hl Hr. 40 | 41 | Tactic Notation "rinversion" hyp(H) := 42 | let Hl := fresh H "l" in 43 | let Hr := fresh H "r" in 44 | rinversion_tac H Hl Hr. 45 | 46 | Tactic Notation "rinversion" constr(H) := 47 | let Hl := fresh "Hl" in 48 | let Hr := fresh "Hr" in 49 | rinversion_tac H Hl Hr. 50 | 51 | (** ** Introducing the hypothesis to be destructed *) 52 | 53 | (** The goal of [rdestruct] is usually to make progress on a goal 54 | relating two [match] constructions, by breaking down the terms being 55 | matched. The instances below will attempt to prove automatically 56 | that these terms are related, then destruct the resulting hypothesis 57 | in order to break them down and reduce the [match]es. 58 | 59 | If proving the destructed hypothesis automatically is not possible, 60 | the user can use the following tactic to introduce it manually. *) 61 | 62 | Ltac rdestruct_assert := 63 | lazymatch goal with 64 | | |- _ (match ?m with _ => _ end) (match ?n with _ => _ end) => 65 | let Tm := type of m in 66 | let Tn := type of n in 67 | let R := fresh "R" in 68 | evar (R: rel Tm Tn); 69 | assert (R m n); subst R 70 | end. 71 | 72 | (** ** [RDestruct] steps *) 73 | 74 | (** Sometimes we want to remember what the destructed terms were. 75 | To make this possible, we use the following relation to wrap the 76 | subgoals generated by the [RDestruct] instance. *) 77 | 78 | Definition rdestruct_result {A B} m n (Q: rel A B): rel A B := 79 | fun x y => m = x /\ n = y -> Q x y. 80 | 81 | (** To use [RDestruct] instances to reduce a goal involving pattern 82 | matching [G] := [_ (match m with _ end) (match n with _ end)], we 83 | need to establish that [m] and [n] are related by some relation [R], 84 | then locate an instance of [RDestruct] that corresponds to [R]. It 85 | is essential that this happens in one step. At some point, I tried 86 | to split the process in two different [RStep]s, so that the user 87 | could control the resolution of the [?R m n] subgoal. However, in 88 | situation where [RDestruct] is not the right strategy, this may 89 | push a [delayed rauto] into a dead end. Thankfully, now we can use 90 | the [RAuto] typeclass to solve the [?R m n] subgoal in one swoop. *) 91 | 92 | Lemma rdestruct_rstep {A B} m n (R: rel A B) P (Q: rel _ _): 93 | RAuto (R m n) -> 94 | RDestruct R P -> 95 | P (rdestruct_result m n Q) -> 96 | Q m n. 97 | Proof. 98 | intros Hmn HR H. 99 | firstorder. 100 | Qed. 101 | 102 | Ltac use_rdestruct_rstep m n := 103 | let H := fresh in 104 | intro H; 105 | pattern m, n; 106 | eapply (rdestruct_rstep m n); 107 | [ .. | eexact H]. 108 | 109 | Global Hint Extern 40 (RStep _ (_ (match ?m with _=>_ end) (match ?n with _=>_ end))) => 110 | use_rdestruct_rstep m n : typeclass_instances. 111 | 112 | (** ** Choosing to discard or keep equations *) 113 | 114 | (** The [RStep] above will leave us with a bunch of [rdestruct_result] 115 | subgoals. The next [RStep] will unpack them, and either discard or 116 | keep equations. 117 | 118 | By default, we discard equations. This is to prevent them from 119 | cluttering the context, possibly interfering with further steps, and 120 | introducing unwanted dependencies. However, the user can use the 121 | [rdestruct_remember] and [rdestruct_forget] tactics below to control 122 | this behavior. *) 123 | 124 | CoInductive rdestruct_remember := rdestruct_remember_intro. 125 | 126 | Ltac rdestruct_remember := 127 | lazymatch goal with 128 | | _ : rdestruct_remember |- _ => 129 | idtac 130 | | _ => 131 | let H := fresh "Hrdestruct" in 132 | pose proof rdestruct_remember_intro as H 133 | end. 134 | 135 | Ltac rdestruct_forget := 136 | lazymatch goal with 137 | | H : rdestruct_remember |- _ => 138 | clear H 139 | | _ => 140 | idtac 141 | end. 142 | 143 | (** The following [RStep] instance will keep or discard the equations 144 | in [rdestruct_result] based on the user's choice. *) 145 | 146 | Lemma rdestruct_forget_rintro {A B} m n (Q: rel A B) x y: 147 | RIntro (Q x y) (rdestruct_result m n Q) x y. 148 | Proof. 149 | firstorder. 150 | Qed. 151 | 152 | Lemma rdestruct_remember_rintro {A B} m n (Q: rel A B) x y: 153 | RIntro (m = x -> n = y -> Q x y) (rdestruct_result m n Q) x y. 154 | Proof. 155 | firstorder. 156 | Qed. 157 | 158 | Ltac rdestruct_result_rintro := 159 | lazymatch goal with 160 | | _ : rdestruct_remember |- _ => 161 | eapply rdestruct_remember_rintro 162 | | _ => 163 | eapply rdestruct_forget_rintro 164 | end. 165 | 166 | Global Hint Extern 100 (RIntro _ (rdestruct_result _ _ _) _ _) => 167 | rdestruct_result_rintro : typeclass_instances. 168 | 169 | (** ** Default instance *) 170 | 171 | (** We provide a default instance of [RDestruct] by reifying the 172 | behavior of the [destruct] tactic on the relational hypothesis that 173 | we're trying to destruct. *) 174 | 175 | Ltac default_rdestruct := 176 | let m := fresh "m" in 177 | let n := fresh "n" in 178 | let Hmn := fresh "H" m n in 179 | let P := fresh "P" in 180 | let H := fresh in 181 | intros m n Hmn P H; 182 | revert m n Hmn; 183 | delayed_conjunction (intros m n Hmn; destruct Hmn; delay); 184 | pattern P; 185 | eexact H. 186 | 187 | Global Hint Extern 100 (RDestruct _ _) => 188 | default_rdestruct : typeclass_instances. 189 | 190 | (** In the special case where the terms matched on the left- and 191 | right-hand sides are identical, we want to destruct that term 192 | instead. We accomplish this by introducing a special instance of 193 | [RDestruct] for the relation [eq]. *) 194 | 195 | Ltac eq_rdestruct := 196 | let m := fresh "m" in 197 | let n := fresh "n" in 198 | let Hmn := fresh "H" m n in 199 | let P := fresh "P" in 200 | let H := fresh in 201 | intros m n Hmn P H; 202 | revert m n Hmn; 203 | delayed_conjunction (intros m n Hmn; destruct Hmn; destruct m; delay); 204 | pattern P; 205 | eexact H. 206 | 207 | Global Hint Extern 99 (RDestruct eq _) => 208 | eq_rdestruct : typeclass_instances. 209 | -------------------------------------------------------------------------------- /RelClasses.v: -------------------------------------------------------------------------------- 1 | Require Export Coq.Classes.RelationClasses. 2 | Require Export RelDefinitions. 3 | 4 | 5 | (** * More classes for relation properties *) 6 | 7 | (** In the following we introduce more typeclasses for characterizing 8 | properties of relations, in the spirit of the [RelationClasses] 9 | standard library. *) 10 | 11 | (** ** Coreflexive relations *) 12 | 13 | Class Coreflexive {A} (R: relation A) := 14 | coreflexivity: forall x y, R x y -> x = y. 15 | 16 | Global Instance eq_corefl {A}: 17 | Coreflexive (@eq A). 18 | Proof. 19 | firstorder. 20 | Qed. 21 | 22 | Global Instance subrel_eq {A} (R: relation A): 23 | Coreflexive R -> 24 | Related R eq subrel. 25 | Proof. 26 | firstorder. 27 | Qed. 28 | 29 | (** ** Composition *) 30 | 31 | Class RCompose {A B C} (RAB : rel A B) (RBC : rel B C) (RAC : rel A C) := 32 | rcompose : forall x y z, RAB x y -> RBC y z -> RAC x z. 33 | 34 | Ltac rcompose b := 35 | lazymatch goal with 36 | | |- ?R ?a ?c => 37 | apply (rcompose a b c) 38 | end. 39 | 40 | Ltac ercompose := 41 | eapply rcompose. 42 | 43 | (** [Transitive] is a special case of [RCompose]. The following makes 44 | it possible for the [transitivity] tactic to use [RCompose] 45 | instances. *) 46 | 47 | Global Instance rcompose_transitive {A} (R : relation A) : 48 | RCompose R R R -> Transitive R. 49 | Proof. 50 | firstorder. 51 | Qed. 52 | 53 | (** Conversely, we will want to use declared [Transitive] instances to 54 | satisfy [RCompose] queries. To avoid loops in the resolution process 55 | this is only declared as an immediate hint --- we expect all derived 56 | instances to be formulated in terms of [RCompose] rather than 57 | [Transitive]. *) 58 | 59 | Lemma transitive_rcompose `(Transitive) : 60 | RCompose R R R. 61 | Proof. 62 | assumption. 63 | Qed. 64 | 65 | Global Hint Immediate transitive_rcompose : typeclass_instances. 66 | 67 | (** ** Decomposition *) 68 | 69 | (** The converse of [RCompose] is the following "decomposition" property. *) 70 | 71 | Class RDecompose {A B C} (RAB : rel A B) (RBC : rel B C) (RAC : rel A C) := 72 | rdecompose : forall x z, RAC x z -> exists y, RAB x y /\ RBC y z. 73 | 74 | Tactic Notation "rdecompose" constr(H) "as" simple_intropattern(p) := 75 | lazymatch type of H with 76 | | ?R ?a ?b => 77 | destruct (rdecompose a b H) as p 78 | | _ => 79 | fail "Not an applied relation" 80 | end. 81 | 82 | Tactic Notation "rdecompose" hyp(H) "as" simple_intropattern(p) := 83 | lazymatch type of H with 84 | | ?R ?a ?b => 85 | apply rdecompose in H; destruct H as p 86 | | _ => 87 | fail "Not an applied relation" 88 | end. 89 | 90 | Tactic Notation "rdecompose" constr(H) := 91 | rdecompose H as (? & ? & ?). 92 | 93 | Tactic Notation "rdecompose" hyp(H) := 94 | rdecompose H as (? & ? & ?). 95 | 96 | 97 | (** * Instances for core relators *) 98 | 99 | (** ** [arrow_rel] *) 100 | 101 | Global Instance arrow_refl {A B} (RA : relation A) (RB : relation B) : 102 | Coreflexive RA -> 103 | Reflexive RB -> 104 | Reflexive (RA ++> RB). 105 | Proof. 106 | intros HA HB f a b Hab. 107 | apply coreflexivity in Hab. subst. 108 | reflexivity. 109 | Qed. 110 | 111 | Global Instance arrow_corefl {A B} (RA : relation A) (RB : relation B) : 112 | Reflexive RA -> 113 | Coreflexive RB -> 114 | Coreflexive (RA ++> RB). 115 | Proof. 116 | (* This requires an extensionality axiom *) 117 | Abort. 118 | 119 | Section ARROW_REL_COMPOSE. 120 | Context {A1 A2 A3} (RA12 : rel A1 A2) (RA23 : rel A2 A3) (RA13 : rel A1 A3). 121 | Context {B1 B2 B3} (RB12 : rel B1 B2) (RB23 : rel B2 B3) (RB13 : rel B1 B3). 122 | 123 | Global Instance arrow_rcompose : 124 | RDecompose RA12 RA23 RA13 -> 125 | RCompose RB12 RB23 RB13 -> 126 | RCompose (RA12 ++> RB12) (RA23 ++> RB23) (RA13 ++> RB13). 127 | Proof. 128 | intros HA HB f g h Hfg Hgh a1 a3 Ha. 129 | rdecompose Ha as (a2 & Ha12 & Ha23). 130 | firstorder. 131 | Qed. 132 | 133 | Global Instance arrow_rdecompose : 134 | RCompose RA12 RA23 RA13 -> 135 | RDecompose RB12 RB23 RB13 -> 136 | RDecompose (RA12 ++> RB12) (RA23 ++> RB23) (RA13 ++> RB13). 137 | Proof. 138 | (* This requires a choice axiom. *) 139 | Abort. 140 | End ARROW_REL_COMPOSE. 141 | -------------------------------------------------------------------------------- /RelDefinitions.v: -------------------------------------------------------------------------------- 1 | Require Export Coq.Program.Basics. 2 | Require Export Coq.Relations.Relation_Definitions. 3 | Require Export Coq.Classes.Morphisms. 4 | Require Setoid. 5 | Require Export Delay. 6 | 7 | (** * Prerequisites *) 8 | 9 | (** Some instances that would normally cause loops can be used 10 | nontheless if we insist that some parameters cannot be existential 11 | variables. One way to do this is to use this guard class, similar in 12 | spirit to [Unconvertible]. *) 13 | 14 | Class NotEvar {A} (x: A). 15 | 16 | Global Hint Extern 1 (NotEvar ?x) => 17 | not_evar x; constructor : typeclass_instances. 18 | 19 | (** This version of [Unconvertible] does not assume that [a] and [b] 20 | have convertible types. *) 21 | 22 | Class Unconvertible {A B} (a: A) (b: B) := unconvertible : unit. 23 | 24 | Ltac unconvertible a b := 25 | first 26 | [ unify a b with typeclass_instances; fail 1 27 | | exact tt ]. 28 | 29 | Global Hint Extern 1 (Unconvertible ?a ?b) => 30 | unconvertible a b : typeclass_instances. 31 | 32 | (** Sometimes we may want to introduce an auxiliary variable to help 33 | with unification. *) 34 | 35 | Class Convertible {A} (x y: A) := 36 | convertible: x = y. 37 | 38 | Global Hint Extern 1 (Convertible ?x ?y) => 39 | eapply eq_refl : typeclass_instances. 40 | 41 | (** The following class can be used to inhibit backtracking. *) 42 | 43 | Class Once P := once : P. 44 | 45 | Global Hint Extern 1 (Once ?P) => 46 | red; once typeclasses eauto : typeclass_instances. 47 | 48 | 49 | (** * Relations *) 50 | 51 | (** The constructions in [Coq.Relations.Relation_Definitions] are only 52 | concerned with relations within a single type, so that [relation A] 53 | is defined as [A -> A -> Prop]. We will need more general 54 | relations, and so I define [rel A B] as [A -> B -> Prop]. *) 55 | 56 | Definition rel (A1 A2: Type) := A1 -> A2 -> Prop. 57 | 58 | Declare Scope rel_scope. 59 | Delimit Scope rel_scope with rel. 60 | Bind Scope rel_scope with rel. 61 | 62 | (** Make sure that the existing definitions based on 63 | [Relation_Definitions.relation] use our [rel_scope] as well. *) 64 | 65 | Bind Scope rel_scope with Relation_Definitions.relation. 66 | 67 | (** ** Proof step *) 68 | 69 | (** This is a catch-all class for any applicable strategy allowing us 70 | to make progress towards solving a relational goal. The proposition 71 | [P] should encode the new subgoals, in the format expected by 72 | [Delay.split_conjunction], and is decomposed accordingly in the 73 | [rstep] tactic. 74 | 75 | At the moment, the following priorities are used for our different 76 | [RStep] instances. Generally speaking, the most specific, quick 77 | tactics should be registered with higher priority (lower numbers), 78 | and the more general, slow tactics that are likely to instantiate 79 | evars incorrectly or spawn them too early should have lower priority 80 | (higher numbers): 81 | 82 | - 10 [Reflexivity] 83 | - 20 [RIntro] 84 | - 30 [preorder] 85 | - 40 [RDestruct] 86 | - 50 [Monotonicity] (includes [Reflexivity] -- we may want to split) 87 | - 70 [RExists] *) 88 | 89 | Class RStep (P Q: Prop) := 90 | rstep: P -> Q. 91 | 92 | Ltac rstep := 93 | lazymatch goal with 94 | | |- ?Q => 95 | apply (rstep (Q := Q)); 96 | Delay.split_conjunction 97 | end. 98 | 99 | (** ** Proof automation *) 100 | 101 | (** The following class solves the goal [Q] automatically. *) 102 | 103 | Class RAuto (Q: Prop) := 104 | rauto : Q. 105 | 106 | Ltac rauto := 107 | lazymatch goal with 108 | | |- ?Q => 109 | apply (rauto (Q := Q)); 110 | Delay.split_conjunction 111 | end. 112 | 113 | (** After applying each step, we need to decompose the premise into 114 | individual subgoals, wrapping each one into a new [RAuto] goal so 115 | that the process continues. 116 | 117 | Note the use of [Once] below: while choosing a step to apply next 118 | can involve some backtracking, once a step has been chosen [RAuto] 119 | never backtracks. This avoids exponential blow up in the search 120 | space, so that [RAuto] remains efficient even in case of failure. 121 | This encourages proper hygiene when declaring instances, since 122 | extraneous or too broadly applicable instance will cause failures 123 | (rather than silently add weight to the system). 124 | This also enables the user to investigate failures by stepping 125 | through the process using the [rstep] tactic. *) 126 | 127 | Class RAutoSubgoals (P: Prop) := 128 | rauto_subgoals : P. 129 | 130 | Global Instance rauto_rstep P Q: 131 | Once (RStep P Q) -> 132 | RAutoSubgoals P -> 133 | RAuto Q. 134 | Proof. 135 | firstorder. 136 | Qed. 137 | 138 | Ltac rauto_split := 139 | red; 140 | Delay.split_conjunction; 141 | lazymatch goal with 142 | | |- ?Q => change (RAuto Q) 143 | end. 144 | 145 | Global Hint Extern 1 (RAutoSubgoals _) => 146 | rauto_split : typeclass_instances. 147 | 148 | (** If [rauto] is run under the [delayed] tactical and we don't know 149 | how to make progress, bail out. *) 150 | 151 | Global Hint Extern 1000 (RAuto _) => 152 | red; solve [ delay ] : typeclass_instances. 153 | 154 | (** ** Reflexivity *) 155 | 156 | (** This is an early check to solve easy goals where a reflexive 157 | relation is applied to two identical arguments, however we are 158 | careful not to cause any instantiation of existential variables. 159 | The [Monotonicity] strategy provides a more general way to use 160 | reflexivity. *) 161 | 162 | Ltac no_evars t := 163 | lazymatch t with 164 | | ?f ?x => no_evars f; no_evars x 165 | | ?m => not_evar m 166 | end. 167 | 168 | Lemma reflexivity_rstep {A} (R: rel A A) (x: A) : 169 | Reflexive R -> 170 | RStep True (R x x). 171 | Proof. 172 | firstorder. 173 | Qed. 174 | 175 | Global Hint Extern 10 (RStep _ (?R ?x ?x)) => 176 | no_evars R; eapply reflexivity_rstep : typeclass_instances. 177 | 178 | (** ** Introduction rules *) 179 | 180 | (** In effect, [RIntro] is pretty much identical to [RStep], but we 181 | like to separate introduction rules and the [rintro] tactic. *) 182 | 183 | Class RIntro {A B} (P: Prop) (R: rel A B) (m: A) (n: B): Prop := 184 | rintro: P -> R m n. 185 | 186 | Arguments RIntro {A%type B%type} P%type R%rel m n. 187 | 188 | Ltac rintro := 189 | lazymatch goal with 190 | | |- ?R ?m ?n => 191 | apply (rintro (R:=R) (m:=m) (n:=n)); 192 | Delay.split_conjunction 193 | end. 194 | 195 | Global Instance rintro_rstep: 196 | forall `(RIntro), RStep P (R m n) | 20. 197 | Proof. 198 | firstorder. 199 | Qed. 200 | 201 | (** Most introduction rules are entierly reversible. For instance, 202 | suppose we use the introduction rule for [++>] on a goal of the form 203 | [(R1 ++> R2) f g], to obtain [Hxy: R1 x y |- R2 (f x) (g y)]. If at 204 | a later stage we manage to prove our old goal [Hfg: (R1 ++> R2) f g], 205 | we can always use the elimination rule for [++>] in conjunction with 206 | the two hypotheses to prove [R2 (f x) (g y)]. 207 | 208 | On the other hand, whenever a new existential variable is involved, 209 | this reversibility is lost: the time of introduction of the evar 210 | determines what a valid instantiation is, and there is no way to go 211 | back if we want to use components introduced later, say by 212 | destructing one of the hypotheses. 213 | 214 | For this reason, we want such introduction rules to be used only as 215 | a last resort, and segregate them as instances of the following 216 | class rather than [RIntro]. Moreover, to make sure that we don't 217 | leave the user in a dead-end, we only use it if we can immediately 218 | solve the resulting subgoal. *) 219 | 220 | Class RExists {A B} (P: Prop) (R: rel A B) (m: A) (n: B): Prop := 221 | rexists: P -> R m n. 222 | 223 | Arguments RExists {A%type B%type} P%type R%rel m n. 224 | 225 | Ltac reexists := 226 | lazymatch goal with 227 | | |- ?R ?m ?n => 228 | apply (rexists (R:=R) (m:=m) (n:=n)); 229 | Delay.split_conjunction 230 | end. 231 | 232 | Global Instance rexists_rstep {A B} P R (m:A) (n:B): 233 | RExists P R m n -> 234 | NonDelayed (RAutoSubgoals P) -> 235 | RStep True (R m n) | 70. 236 | Proof. 237 | firstorder. 238 | Qed. 239 | 240 | (** ** Using relations *) 241 | 242 | (** As we extend our relations language with new relators, we need to 243 | be able to extend the ways in which corresponding relational 244 | properties can be applied to a given goal. The following type class 245 | expresses that the relational property [R m n] can be applied to a 246 | goal of type [Q], generating the subgoal [P]. *) 247 | 248 | Class RElim {A B} (R: rel A B) (m: A) (n: B) (P Q: Prop): Prop := 249 | relim: R m n -> P -> Q. 250 | 251 | Arguments RElim {A%type B%type} R%rel m n P%type Q%type. 252 | 253 | Ltac relim H := 254 | lazymatch goal with 255 | | |- ?Q => 256 | apply (relim (Q:=Q) H) 257 | (* ^ XXX split_conjunction? *) 258 | end. 259 | 260 | (** The resolution process is directed by the syntax of [R]. We define 261 | an instance for each function relator. The base case simply uses the 262 | relational property as-is. *) 263 | 264 | Global Instance relim_base {A B} (R: rel A B) m n: 265 | RElim R m n True (R m n) | 10. 266 | Proof. 267 | firstorder. 268 | Qed. 269 | 270 | (** ** Destructing relational hypotheses *) 271 | 272 | (** To make progress when the goal relates two pattern matching 273 | constructions, we need to show that the two matched terms are 274 | related, then destruct that hypothesis in a such a way that the two 275 | terms reduce to constructors. 276 | 277 | For most relators of inductive types, the constructors of the 278 | relator will simply follow the constructors of the corresponding 279 | type, so that destructing the relational hypothesis in the usual way 280 | will produce the desired outcome. However, sometimes it is more 281 | convenient to define such relators in a different way (see for 282 | instance [prod_rel]). In that case, we can use the following 283 | typeclass to specify an alternative way to destruct corresponding 284 | hypotheses. 285 | 286 | An instance of [RDestruct] is somewhat similar to a stylized 287 | induction principle. [T] expands to a conjunction of subgoals in the 288 | format expected by [Delay.split_conjunction]. For instance, the 289 | induction principle for [sum_rel] is: 290 | << 291 | sum_rel_ind: 292 | forall ..., 293 | (forall a1 a2, RA a1 a2 -> P (inl a1) (inl a2)) -> 294 | (forall b1 b2, RB b1 b2 -> P (inr b1) (inr b2)) -> 295 | (forall p1 p2, (RA + RB) p1 p2 -> P p1 p2) 296 | >> 297 | A corresponding instance of [RDestruct] would be: 298 | << 299 | sum_rdestruct: 300 | RDestruct 301 | (sum_rel RA RB) 302 | (fun P => 303 | (forall a1 a2, RA a1 a2 -> P (inl a1) (inl a2)) /\ 304 | (forall b1 b2, RB b1 b2 -> P (inr b1) (inr b2))) 305 | >> 306 | In the case of [sum_rel] however, which is defined as an inductive 307 | type with similar structure to [sum], we can rely on the default 308 | instance of [RDestruct], which simply uses the [destruct] tactic. *) 309 | 310 | Class RDestruct {A B: Type} (R: rel A B) (T: rel A B -> Prop) := 311 | rdestruct m n: R m n -> forall P, T P -> P m n. 312 | 313 | (** See the [RDestruct] library for the corresponding instance of 314 | [RStep], the default instance of [RDestruct], and a way to control 315 | whether or not we should keep equations for the destructed terms. *) 316 | 317 | (** ** Monotonicity properties *) 318 | 319 | (** We use the class [Related] for the user to declare monotonicity 320 | properties. This is a generalization of [Morphisms.Proper] from the 321 | Coq standard library: although we expect that most of the time the 322 | left- and right-hand side terms will be identical, they could be 323 | slightly different partial applications of the same function. 324 | 325 | Usually the differing arguments will be implicit, so that the user 326 | can still rely on the [Monotonic] notation below. Occasionally, you 327 | may need to spell out the two different terms and use the actual 328 | class instead. 329 | 330 | Note that the argument order below eschews the precedent of [Proper], 331 | which has the relation first, followed by the proper element. This is 332 | deliberate: we want the type parameters [A] and [B] to be unified in 333 | priority against the type of [f] and [g], rather than that of [R]. 334 | In particular, the relator [forall_rel] yields an eta-expanded type 335 | of the form [(fun x => T x) x] for its arguments. When [A] and [B] 336 | take this peculiar form, instances declared using [forall_rel] 337 | become unusable (I assume because no second-order unification is 338 | performed when looking up the typeclass instance database). By 339 | contrast the type of [f] and [g] is much more likely to be in a 340 | "reasonable" form. *) 341 | 342 | Class Related {A B} (f: A) (g: B) (R: rel A B) := 343 | related: R f g. 344 | 345 | Arguments Related {A%type B%type} _ _ R%rel. 346 | 347 | Notation "'@' 'Monotonic' T m R" := (@Related T T m m R%rel) 348 | (at level 10, T at next level, R at next level, m at next level). 349 | 350 | Notation Monotonic m R := (Related m m R%rel). 351 | 352 | (** We provide a [RStep] instance for unfolding [Related]. *) 353 | 354 | Lemma unfold_monotonic_rstep {A B} (R: rel A B) m n: 355 | RStep (R m n) (Related m n R). 356 | Proof. 357 | firstorder. 358 | Qed. 359 | 360 | Global Hint Extern 1 (RStep _ (Related _ _ _)) => 361 | eapply unfold_monotonic_rstep : typeclass_instances. 362 | 363 | (** ** Order on relations *) 364 | 365 | (** This is our generalization of [subrelation]. Like the original it 366 | constitutes a preorder, and the union and intersection of relations 367 | are the corresponding join and meet. *) 368 | 369 | Definition subrel {A B}: rel (rel A B) (rel A B) := 370 | fun R1 R2 => forall x y, R1 x y -> R2 x y. 371 | 372 | Arguments subrel {A%type B%type} R1%rel R2%rel. 373 | 374 | Global Instance subrel_preorder A B: 375 | @PreOrder (rel A B) subrel. 376 | Proof. 377 | split; firstorder. 378 | Qed. 379 | 380 | Global Instance eq_subrel {A} (R: rel A A): 381 | Reflexive R -> 382 | Related eq R subrel. 383 | Proof. 384 | intros HR x y H. 385 | subst. 386 | reflexivity. 387 | Qed. 388 | 389 | Instance subrel_impl_relim {A B} (R1 R2 : rel A B) x1 x2 y1 y2 P Q: 390 | RElim impl (R1 x1 y1) (R2 x2 y2) P Q -> 391 | RElim subrel R1 R2 (x1 = x2 /\ y1 = y2 /\ P) Q. 392 | Proof. 393 | cbv. 394 | firstorder. 395 | subst. 396 | eauto. 397 | Qed. 398 | 399 | 400 | (** * Core relators *) 401 | 402 | (** First, we introduce the core relators necessary for everything 403 | else, namely those for function types. The next section provides a 404 | more comprehensive library which covers many of the basic inductive 405 | type constructors as well. *) 406 | 407 | (** *** Non-dependent function types *) 408 | 409 | (** First, I define relators for non-dependent functions. This 410 | generalizes [respectful]. *) 411 | 412 | Definition arrow_rel {A1 A2 B1 B2}: 413 | rel A1 A2 -> rel B1 B2 -> rel (A1 -> B1) (A2 -> B2) := 414 | fun RA RB f g => forall x y, RA x y -> RB (f x) (g y). 415 | 416 | Arguments arrow_rel {A1%type A2%type B1%type B2%type} RA%rel RB%rel _ _. 417 | 418 | Notation "RA ==> RB" := (arrow_rel RA RB) 419 | (at level 55, right associativity) : rel_scope. 420 | 421 | Notation "RA ++> RB" := (arrow_rel RA RB) 422 | (at level 55, right associativity) : rel_scope. 423 | 424 | Notation "RA --> RB" := (arrow_rel (flip RA) RB) 425 | (at level 55, right associativity) : rel_scope. 426 | 427 | Global Instance arrow_subrel {A1 A2 B1 B2}: 428 | Monotonic (@arrow_rel A1 A2 B1 B2) (subrel --> subrel ++> subrel). 429 | Proof. 430 | firstorder. 431 | Qed. 432 | 433 | Global Instance arrow_subrel_params: 434 | Params (@arrow_rel) 4 := { }. 435 | 436 | Lemma arrow_rintro {A1 A2 B1 B2} (RA: rel A1 A2) (RB: rel B1 B2) f g: 437 | RIntro (forall x y, RA x y -> RB (f x) (g y)) (RA ++> RB) f g. 438 | Proof. 439 | firstorder. 440 | Qed. 441 | 442 | Global Hint Extern 0 (RIntro _ (_ ++> _) _ _) => 443 | eapply arrow_rintro : typeclass_instances. 444 | 445 | Lemma arrow_relim {A1 A2 B1 B2} RA RB f g m n P Q: 446 | @RElim B1 B2 RB (f m) (g n) P Q -> 447 | @RElim (A1 -> B1) (A2 -> B2) (RA ++> RB) f g (RA m n /\ P) Q. 448 | Proof. 449 | firstorder. 450 | Qed. 451 | 452 | Global Hint Extern 1 (RElim (_ ++> _) _ _ _ _) => 453 | eapply arrow_relim : typeclass_instances. 454 | 455 | (** *** Dependent products *) 456 | 457 | (** Now we consider the dependent case. The definition of [forall_rel] 458 | is somewhat involved, but you can think of relating [f] and [g] in 459 | the context of a structure-preserving transformation from a quiver 460 | ([V], [E]) to the quiver ([Type], [rel]). Like a functor, it has two 461 | components: [FV] maps nodes to types, and [FE] maps edges to 462 | relations. Then, [forall_rel FE f g] states that given an edge 463 | [(e : E v1 v2)], the images [f v1] and [g v2] are related by the 464 | corresponding relation [FE v1 v2 e]. We will write [forall_rel FE f g] 465 | as [(forallr e @ v1 v2 : E, FE[v1,v2,e]) f g]. Notice that this notation 466 | binds [v1] and [v2] as well as [e]. 467 | 468 | If that makes no sense, you can think of specific source quivers. So 469 | for instance, oftentimes we will want to use ([Type], [rel]) as the 470 | source quiver too. This corresponds to parametric polymorphism. The 471 | type of [Some] is [forall A : Type, A -> option A]; the corresponding 472 | logical relation is [forallr R @ A1 A2 : rel, R ++> option_rel R]. Stating 473 | that [Monotonic Some (foralr R @ A1 A2 : rel, R ++> option_rel R)] means that, 474 | given any relation [R] and elements [x1] and [x2], if [R] relates 475 | [x1] and [x2], then [option_rel R] will relate [Some x1] and [Some x2]. 476 | 477 | Another example from [liblayers] is the quiver of our data-indexed 478 | simulation relations [simrel : layerdata -> layerdata -> Type]. 479 | Here the structure-preserving transformations are our simulation 480 | relation diagrams, which have types such as 481 | [lsim : forall D1 D2, simrel D1 D2 -> rel (layer D1) (layer D2)] or 482 | [psim : forall {D1 D2}, simrel D1 D2 -> rel (primsem D1) (primsem D2)]. 483 | Then, the monotonicity of a data-indexed function — 484 | say, [foo: forall D : layerdata, layer D -> primsem D] — 485 | can be expressed as 486 | [Monotonic foo (forall R @ D1 D2 : simrel, siml D1 D2 R ++> simp D1 D2 R) foo]. 487 | 488 | This definition is the same as [respectful_hetero]. *) 489 | 490 | Definition forall_rel {V1 V2} {E: V1->V2->Type} {FV1: V1->Type} {FV2: V2->Type}: 491 | (forall v1 v2, E v1 v2 -> rel (FV1 v1) (FV2 v2)) -> 492 | rel (forall v1, FV1 v1) (forall v2, FV2 v2) := 493 | fun FE f g => 494 | forall v1 v2 (e: E v1 v2), FE v1 v2 e (f v1) (g v2). 495 | 496 | Arguments forall_rel {V1%type V2%type E%type FV1%type FV2%type} FE%rel _ _. 497 | 498 | Notation "'forallr' e @ v1 v2 : E , R" := 499 | (forall_rel (E := E) (fun v1 v2 e => R)) 500 | (at level 200, e ident, v1 ident, v2 ident, right associativity) : rel_scope. 501 | 502 | Notation "'forallr' e @ v1 v2 , R" := 503 | (forall_rel (fun v1 v2 e => R)) 504 | (at level 200, e ident, v1 ident, v2 ident, right associativity) : rel_scope. 505 | 506 | Notation "'forallr' e : E , R" := 507 | (forall_rel (E := E) (fun _ _ e => R)) 508 | (at level 200, e ident, right associativity) : rel_scope. 509 | 510 | Notation "'forallr' e , R" := 511 | (forall_rel (fun _ _ e => R)) 512 | (at level 200, e ident, right associativity) : rel_scope. 513 | 514 | Lemma forall_rintro {V1 V2 E F1 F2} (FE: forall x y, _ -> rel _ _) f g: 515 | RIntro 516 | (forall u v e, FE u v e (f u) (g v)) 517 | (@forall_rel V1 V2 E F1 F2 FE) f g. 518 | Proof. 519 | firstorder. 520 | Qed. 521 | 522 | Global Hint Extern 0 (RIntro _ (forall_rel _) _ _) => 523 | eapply forall_rintro : typeclass_instances. 524 | 525 | Lemma forall_relim {V1 V2 E FV1 FV2} R f g v1 v2 e P Q: 526 | RElim (R v1 v2 e) (f v1) (g v2) P Q -> 527 | RElim (@forall_rel V1 V2 E FV1 FV2 R) f g P Q. 528 | Proof. 529 | firstorder. 530 | Qed. 531 | 532 | Global Hint Extern 1 (RElim (forall_rel _) _ _ _ _) => 533 | eapply forall_relim : typeclass_instances. 534 | 535 | (** ** Inverse relation *) 536 | 537 | (** We use [flip] as our inversion relator. To this end we 538 | characterize its variance and introduce an [RElim] rule. *) 539 | 540 | Global Instance flip_subrel {A B}: 541 | Monotonic (@flip A B Prop) (subrel ++> subrel). 542 | Proof. 543 | firstorder. 544 | Qed. 545 | 546 | Global Instance flip_subrel_params: 547 | Params (@flip) 3 := { }. 548 | 549 | Lemma flip_rintro {A B} (R: rel A B) m n: 550 | RIntro (R n m) (flip R) m n. 551 | Proof. 552 | firstorder. 553 | Qed. 554 | 555 | Global Hint Extern 1 (RIntro _ (flip _) _ _) => 556 | eapply flip_rintro : typeclass_instances. 557 | 558 | Lemma flip_relim {A B} (R: rel A B) m n P Q: 559 | RElim R n m P Q -> 560 | RElim (flip R) m n P Q. 561 | Proof. 562 | firstorder. 563 | Qed. 564 | 565 | Global Hint Extern 1 (RElim (flip _) _ _ _ _) => 566 | eapply flip_relim : typeclass_instances. 567 | 568 | Lemma flip_rdestruct {A B} (R: rel A B) T: 569 | RDestruct R T -> 570 | RDestruct (flip R) (fun P => T (flip P)). 571 | Proof. 572 | firstorder. 573 | Qed. 574 | 575 | Global Hint Extern 1 (RDestruct (flip _) _) => 576 | eapply flip_rdestruct : typeclass_instances. 577 | 578 | (** When the goal is of the form [?R x y] with [?R] an uninstantiated 579 | existential variable, we must then decide whether to attempt 580 | resolving [?R x y] directly, or try resolving [?R' y x] instead, 581 | with [?R] instantiated to [flip ?R']. However, it is impossible to 582 | know in advance which of these will eventually work, so this is one 583 | situation where backtracking is unavoidable. 584 | 585 | The [RExists] instances below are used to make that decision. 586 | The class [PolarityResolved] can be used as a guard in [RIntro] 587 | instances which are general enough to instantiate such [?R]s — 588 | for example, see the [Monotonicity] library. It ensures that the 589 | specified relation either is not an existential variable, or that 590 | its polarity has been chosen through [direct_rexists] or 591 | [flip_rexists] below. *) 592 | 593 | Class PolarityResolved {A B} (R: rel A B). 594 | 595 | Global Hint Extern 1 (PolarityResolved ?R) => 596 | not_evar R; constructor : typeclass_instances. 597 | 598 | Ltac polarity_unresolved R := 599 | is_evar R; 600 | lazymatch goal with 601 | | H : PolarityResolved R |- _ => fail 602 | | _ => idtac 603 | end. 604 | 605 | Lemma direct_rexists {A B} (R: rel A B) (m: A) (n: B): 606 | RExists (PolarityResolved R -> R m n) R m n. 607 | Proof. 608 | firstorder. 609 | Qed. 610 | 611 | Global Hint Extern 1 (RExists _ ?R _ _) => 612 | polarity_unresolved R; eapply direct_rexists : typeclass_instances. 613 | 614 | Lemma flip_rexists {A B} (R: rel A B) (Rc: rel B A) (m: A) (n: B): 615 | Convertible R (flip Rc) -> 616 | RExists (PolarityResolved Rc -> flip Rc m n) R m n. 617 | Proof. 618 | unfold Convertible. intros; subst. 619 | firstorder. 620 | Qed. 621 | 622 | Global Hint Extern 2 (RExists _ ?R _ _) => 623 | polarity_unresolved R; eapply flip_rexists : typeclass_instances. 624 | -------------------------------------------------------------------------------- /RelOperators.v: -------------------------------------------------------------------------------- 1 | Require Export RelDefinitions. 2 | Require Import RelClasses. 3 | Require Import Relators. 4 | 5 | (** ** Union of relations *) 6 | 7 | Definition rel_union {A B} (R1 R2: rel A B): rel A B := 8 | fun x y => R1 x y \/ R2 x y. 9 | 10 | Arguments rel_union {_ _} R1%rel R2%rel _ _. 11 | 12 | Infix "\/" := rel_union : rel_scope. 13 | 14 | Global Instance rel_union_subrel {A B}: 15 | Monotonic (@rel_union A B) (subrel ++> subrel ++> subrel). 16 | Proof. 17 | clear. 18 | firstorder. 19 | Qed. 20 | 21 | Global Instance rel_union_subrel_params: 22 | Params (@rel_union) 4 := { }. 23 | 24 | (** Since we're forced to commit to a branch, we can't use [RIntro], 25 | but can still provide [RExists] instances. *) 26 | 27 | Lemma rel_union_rexists_l {A B} (R1 R2: rel A B) x y: 28 | RExists (R1 x y) (R1 \/ R2) x y. 29 | Proof. 30 | firstorder. 31 | Qed. 32 | 33 | Global Hint Extern 0 (RExists _ (_ \/ _) _ _) => 34 | eapply rel_union_rexists_l : typeclass_instances. 35 | 36 | Lemma rel_union_rexists_r {A B} (R1 R2: rel A B) x y: 37 | RExists (R2 x y) (R1 \/ R2) x y. 38 | Proof. 39 | firstorder. 40 | Qed. 41 | 42 | Global Hint Extern 0 (RExists _ (_ \/ _) _ _) => 43 | eapply rel_union_rexists_r : typeclass_instances. 44 | 45 | (** More often, we can solve a [rel_union] goal using a monotonicity 46 | property, and the [subrel] instances below. *) 47 | 48 | Lemma rel_union_subrel_rexists_l {A B} (R R1 R2: rel A B): 49 | RExists (subrel R R1) subrel R (R1 \/ R2)%rel. 50 | Proof. 51 | firstorder. 52 | Qed. 53 | 54 | Global Hint Extern 0 (RExists _ subrel _ (_ \/ _)%rel) => 55 | eapply rel_union_subrel_rexists_l : typeclass_instances. 56 | 57 | Lemma rel_union_subrel_rexists_r {A B} (R R1 R2: rel A B): 58 | RExists (subrel R R2) subrel R (R1 \/ R2)%rel. 59 | Proof. 60 | firstorder. 61 | Qed. 62 | 63 | Global Hint Extern 0 (RExists _ subrel _ (_ \/ _)%rel) => 64 | eapply rel_union_subrel_rexists_r : typeclass_instances. 65 | 66 | Lemma rel_union_lub {A B} (R1 R2 R: rel A B): 67 | RIntro (subrel R1 R /\ subrel R2 R) subrel (R1 \/ R2)%rel R. 68 | Proof. 69 | firstorder. 70 | Qed. 71 | 72 | Global Hint Extern 2 (RIntro _ subrel (_ \/ _)%rel _) => 73 | eapply rel_union_lub : typeclass_instances. 74 | 75 | Lemma rel_union_refl_l {A} (R1 R2: rel A A): 76 | Reflexive R1 -> 77 | Reflexive (R1 \/ R2). 78 | Proof. 79 | firstorder. 80 | Qed. 81 | 82 | Global Hint Extern 1 (Reflexive (_ \/ _)) => 83 | eapply rel_union_refl_l : typeclass_instances. 84 | 85 | Lemma rel_union_refl_r {A} (R1 R2: rel A A): 86 | Reflexive R2 -> 87 | Reflexive (R1 \/ R2). 88 | Proof. 89 | firstorder. 90 | Qed. 91 | 92 | Global Hint Extern 1 (Reflexive (_ \/ _)) => 93 | eapply rel_union_refl_r : typeclass_instances. 94 | 95 | Lemma rel_union_corefl {A} (R1 R2: rel A A): 96 | Coreflexive R1 -> 97 | Coreflexive R2 -> 98 | Coreflexive (R1 \/ R2). 99 | Proof. 100 | firstorder. 101 | Qed. 102 | 103 | Global Hint Extern 1 (Coreflexive (_ \/ _)) => 104 | eapply rel_union_corefl : typeclass_instances. 105 | 106 | Lemma rel_union_sym {A} (R1 R2: rel A A): 107 | Symmetric R1 -> 108 | Symmetric R2 -> 109 | Symmetric (R1 \/ R2). 110 | Proof. 111 | firstorder. 112 | Qed. 113 | 114 | Global Hint Extern 1 (Symmetric (_ \/ _)) => 115 | eapply rel_union_sym : typeclass_instances. 116 | 117 | (** ** Intersection of relations *) 118 | 119 | Definition rel_inter {A B} (R1 R2: rel A B): rel A B := 120 | fun x y => R1 x y /\ R2 x y. 121 | 122 | Arguments rel_inter {_ _} R1%rel R2%rel _ _. 123 | 124 | Infix "/\" := rel_inter : rel_scope. 125 | 126 | Global Instance rel_inter_subrel {A B}: 127 | Monotonic (@rel_inter A B) (subrel ++> subrel ++> subrel). 128 | Proof. 129 | clear. 130 | firstorder. 131 | Qed. 132 | 133 | Global Instance rel_inter_params: 134 | Params (@rel_inter) 4 := { }. 135 | 136 | (** In some cases, the relation in the goal will cause existential 137 | variables to be instantiated accordingly; for instance, [rstep] run 138 | on [?R x y |- (R1 /\ R2) x y] will instantiate [?R := R1 /\ R2]. 139 | Because of this, splitting a [rel_inter] goal into two subgoals may 140 | cause a dead-end and we have to use [RExists] for the following 141 | rule. *) 142 | 143 | Lemma rel_inter_rexists {A B} (R1 R2: rel A B) x y: 144 | RExists (R1 x y /\ R2 x y) (R1 /\ R2) x y. 145 | Proof. 146 | firstorder. 147 | Qed. 148 | 149 | Global Hint Extern 0 (RExists _ (_ /\ _) _ _) => 150 | eapply rel_inter_rexists : typeclass_instances. 151 | 152 | Lemma rel_inter_subrel_rexists_l {A B} (R1 R2 R: rel A B): 153 | RExists (subrel R1 R) subrel (R1 /\ R2)%rel R. 154 | Proof. 155 | firstorder. 156 | Qed. 157 | 158 | Global Hint Extern 0 (RExists _ subrel (_ /\ _)%rel _) => 159 | eapply rel_inter_subrel_rexists_l : typeclass_instances. 160 | 161 | Lemma rel_inter_subrel_rexists_r {A B} (R1 R2 R: rel A B): 162 | RExists (subrel R2 R) subrel (R1 /\ R2)%rel R. 163 | Proof. 164 | firstorder. 165 | Qed. 166 | 167 | Global Hint Extern 0 (RExists _ subrel (_ /\ _)%rel _) => 168 | eapply rel_inter_subrel_rexists_r : typeclass_instances. 169 | 170 | Lemma rel_inter_glb {A B} (R R1 R2: rel A B): 171 | RIntro (subrel R R1 /\ subrel R R2) subrel R (R1 /\ R2)%rel. 172 | Proof. 173 | firstorder. 174 | Qed. 175 | 176 | Global Hint Extern 2 (RIntro _ subrel _ (_ /\ _)%rel) => 177 | eapply rel_inter_glb : typeclass_instances. 178 | 179 | Lemma rel_inter_refl {A} (R1 R2: rel A A): 180 | Reflexive R1 -> 181 | Reflexive R2 -> 182 | Reflexive (R1 /\ R2). 183 | Proof. 184 | intros H1 H2. 185 | split; reflexivity. 186 | Qed. 187 | 188 | Global Hint Extern 2 (Reflexive (_ /\ _)) => 189 | eapply rel_inter_refl : typeclass_instances. 190 | 191 | Lemma rel_inter_corefl_l {A} (R1 R2: rel A A): 192 | Coreflexive R1 -> 193 | Coreflexive (R1 /\ R2). 194 | Proof. 195 | firstorder. 196 | Qed. 197 | 198 | Global Hint Extern 1 (Coreflexive (_ /\ _)) => 199 | eapply rel_inter_corefl_l : typeclass_instances. 200 | 201 | Lemma rel_inter_corefl_r {A} (R1 R2: rel A A): 202 | Coreflexive R2 -> 203 | Coreflexive (R1 /\ R2). 204 | Proof. 205 | firstorder. 206 | Qed. 207 | 208 | Global Hint Extern 1 (Coreflexive (_ /\ _)) => 209 | eapply rel_inter_corefl_r : typeclass_instances. 210 | 211 | Lemma rel_inter_trans {A} (R1 R2: rel A A): 212 | Transitive R1 -> 213 | Transitive R2 -> 214 | Transitive (R1 /\ R2). 215 | Proof. 216 | intros H1 H2 x y z [Hxy1 Hxy2] [Hyz1 Hyz2]. 217 | split; etransitivity; eassumption. 218 | Qed. 219 | 220 | Global Hint Extern 2 (Transitive (_ /\ _)) => 221 | eapply rel_inter_trans : typeclass_instances. 222 | 223 | Lemma rel_inter_sym {A} (R1 R2: rel A A): 224 | Symmetric R1 -> 225 | Symmetric R2 -> 226 | Symmetric (R1 /\ R2). 227 | Proof. 228 | intros H1 H2 x y [Hxy1 Hxy2]. 229 | split; symmetry; assumption. 230 | Qed. 231 | 232 | Global Hint Extern 2 (Symmetric (_ /\ _)) => 233 | eapply rel_inter_sym : typeclass_instances. 234 | 235 | Global Instance rel_inter_flip_sym {A} (R: rel A A): 236 | Symmetric (R /\ flip R). 237 | Proof. 238 | intros x y [Hxy Hyx]. 239 | split; assumption. 240 | Qed. 241 | 242 | (** On a related note, a symmetric subrelation of [R] is also a 243 | subrelation of its inverse. *) 244 | 245 | Lemma subrel_sym_flip {A} (R R': relation A): 246 | Symmetric R -> 247 | RStep (subrel R R') (subrel R (flip R')). 248 | Proof. 249 | intros HR H x y Hxy. 250 | symmetry in Hxy. 251 | firstorder. 252 | Qed. 253 | 254 | Global Hint Extern 60 (RStep _ (subrel _ (flip _))) => 255 | eapply subrel_sym_flip : typeclass_instances. 256 | 257 | (** ** Implication *) 258 | 259 | Definition rel_impl {A B} (R1 R2: rel A B): rel A B := 260 | fun x y => R1 x y -> R2 x y. 261 | 262 | Global Instance rel_impl_subrel {A B}: 263 | Monotonic (@rel_impl A B) (subrel --> subrel ++> subrel). 264 | Proof. 265 | firstorder. 266 | Qed. 267 | 268 | Global Instance rel_impl_subrel_params: 269 | Params (@rel_impl) 4 := { }. 270 | 271 | Lemma rel_impl_rintro {A B} (R1 R2: rel A B) x y: 272 | RIntro (R1 x y -> R2 x y) (rel_impl R1 R2) x y. 273 | Proof. 274 | firstorder. 275 | Qed. 276 | 277 | Global Hint Extern 0 (RIntro _ (rel_impl _ _) _ _) => 278 | eapply rel_impl_rintro : typeclass_instances. 279 | 280 | Lemma rel_impl_relim {A B} (R1 R2: rel A B) x y: 281 | RElim (rel_impl R1 R2) x y (R1 x y) (R2 x y). 282 | Proof. 283 | firstorder. 284 | Qed. 285 | 286 | Global Hint Extern 0 (RElim (rel_impl _ _) _ _ _ _) => 287 | eapply rel_impl_relim : typeclass_instances. 288 | 289 | Lemma rel_impl_subrel_codomain {A B} (R1 R2: rel A B): 290 | Related R2 (rel_impl R1 R2) subrel. 291 | Proof. 292 | firstorder. 293 | Qed. 294 | 295 | (** ** The bottom and top relations *) 296 | 297 | Definition rel_bot {A B}: rel A B := 298 | fun x y => False. 299 | 300 | Notation "⊥" := rel_bot : rel_scope. 301 | 302 | Lemma rel_bot_subrel {A B} (R: rel A B): 303 | Related ⊥%rel R subrel. 304 | Proof. 305 | firstorder. 306 | Qed. 307 | 308 | Global Hint Extern 0 (Related ⊥%rel _ _) => 309 | eapply rel_bot_subrel : typeclass_instances. 310 | 311 | Lemma rel_bot_relim {A B} (x: A) (y: B) P: 312 | RElim ⊥ x y True P. 313 | Proof. 314 | firstorder. 315 | Qed. 316 | 317 | Global Hint Extern 0 (RElim ⊥ _ _ _ _) => 318 | eapply rel_bot_relim : typeclass_instances. 319 | 320 | Definition rel_top {A B}: rel A B := 321 | fun x y => True. 322 | 323 | Notation "⊤" := rel_top : rel_scope. 324 | 325 | Lemma rel_top_rintro {A B} (x: A) (y: B): 326 | RIntro True ⊤ x y. 327 | Proof. 328 | firstorder. 329 | Qed. 330 | 331 | Global Hint Extern 0 (RIntro _ ⊤ _ _) => 332 | eapply rel_top_rintro : typeclass_instances. 333 | 334 | Global Instance rel_top_equiv {A}: 335 | @Equivalence A ⊤. 336 | Proof. 337 | repeat constructor. 338 | Qed. 339 | 340 | (** ** Relation equivalence *) 341 | 342 | Definition eqrel {A B}: rel (rel A B) (rel A B) := 343 | (subrel /\ flip subrel)%rel. 344 | 345 | Arguments eqrel {_ _} RA%rel RB%rel. 346 | 347 | Global Instance eqrel_equivalence A B: 348 | Equivalence (@eqrel A B). 349 | Proof. 350 | unfold eqrel. 351 | split; typeclasses eauto. 352 | Qed. 353 | 354 | Global Instance eqrel_subrel A B: 355 | Related (@eqrel A B) (@subrel A B) subrel. 356 | Proof. 357 | firstorder. 358 | Qed. 359 | 360 | (** ** Relation composition *) 361 | 362 | Definition rel_compose {A B C} (RAB: rel A B) (RBC: rel B C): rel A C := 363 | fun x z => exists y, RAB x y /\ RBC y z. 364 | 365 | Global Hint Unfold rel_compose : core. 366 | 367 | Global Instance rel_compose_subrel {A B C}: 368 | Monotonic (@rel_compose A B C) (subrel ++> subrel ++> subrel). 369 | Proof. 370 | firstorder. 371 | Qed. 372 | 373 | Global Instance rel_compose_eqrel {A B C}: 374 | Monotonic (@rel_compose A B C) (eqrel ==> eqrel ==> eqrel). 375 | Proof. 376 | firstorder. 377 | Qed. 378 | 379 | Global Instance rel_compose_params: 380 | Params (@rel_compose) 4 := { }. 381 | 382 | Lemma rel_compose_id_left {A B} (R: rel A B): 383 | eqrel (rel_compose R eq) R. 384 | Proof. 385 | unfold rel_compose. 386 | split; intros x y; firstorder; congruence. 387 | Qed. 388 | 389 | Lemma rel_compose_id_right {A B} (R: rel A B): 390 | eqrel (rel_compose eq R) R. 391 | Proof. 392 | unfold rel_compose. 393 | split; intros x y; firstorder; congruence. 394 | Qed. 395 | 396 | Lemma rel_compose_assoc {A B C D} (RAB: rel A B) (RBC: rel B C) (RCD: rel C D): 397 | eqrel 398 | (rel_compose (rel_compose RAB RBC) RCD) 399 | (rel_compose RAB (rel_compose RBC RCD)). 400 | Proof. 401 | unfold rel_compose. 402 | split; intros x y; firstorder; congruence. 403 | Qed. 404 | 405 | (** Transitivity and decomposition can be expressed in terms of 406 | [rel_compose] in the following way. *) 407 | 408 | Global Instance rel_compose_rcompose {A B C} (RAB : rel A B) (RBC : rel B C) : 409 | RCompose RAB RBC (rel_compose RAB RBC). 410 | Proof. 411 | firstorder. 412 | Qed. 413 | 414 | Global Instance rel_compose_rdecompose {A B C} (RAB : rel A B) (RBC : rel B C) : 415 | RDecompose RAB RBC (rel_compose RAB RBC). 416 | Proof. 417 | firstorder. 418 | Qed. 419 | 420 | Global Instance rcompose_subrel `(RCompose) : 421 | Related (rel_compose RAB RBC) RAC subrel. 422 | Proof. 423 | firstorder. 424 | Qed. 425 | 426 | Global Instance rdecompose_subrel `(RDecompose) : 427 | Related RAC (rel_compose RAB RBC) subrel. 428 | Proof. 429 | firstorder. 430 | Qed. 431 | 432 | 433 | (** ** Pulling a relation along functions *) 434 | 435 | Definition rel_pull {A B A' B'} f g (R: rel A' B'): rel A B := 436 | fun x y => R (f x) (g y). 437 | 438 | (** We use the following notation. Left associativity would make more 439 | sense but we have to match [Coq.Classes.RelationPairs]. *) 440 | 441 | Notation "R @@ ( f , g )" := (rel_pull f g R) 442 | (at level 30, right associativity) : rel_scope. 443 | 444 | Notation "R @@ f" := (rel_pull f f R) 445 | (at level 30, right associativity) : rel_scope. 446 | 447 | Notation "R @@ ( f )" := (rel_pull f f R) 448 | (at level 30, right associativity) : rel_scope. 449 | 450 | Global Instance rel_pull_subrel {A B A' B'} (f: A -> A') (g: B -> B'): 451 | Monotonic (rel_pull f g) (subrel ++> subrel). 452 | Proof. 453 | firstorder. 454 | Qed. 455 | 456 | Global Instance rel_pull_subrel_params: 457 | Params (@rel_pull) 3 := { }. 458 | 459 | (** In the restricted case where [f = g], [rel_pull] preserves many 460 | properties of the underlying relation. *) 461 | 462 | Lemma rel_pull_refl {A B} (f: A -> B) (R: rel B B): 463 | Reflexive R -> 464 | Reflexive (R @@ f). 465 | Proof. 466 | firstorder. 467 | Qed. 468 | 469 | Global Hint Extern 1 (Reflexive (rel_pull _ _ _)) => 470 | eapply rel_pull_refl : typeclass_instances. 471 | 472 | Lemma rel_pull_sym {A B} (f: A -> B) R: 473 | Symmetric R -> 474 | Symmetric (R @@ f). 475 | Proof. 476 | firstorder. 477 | Qed. 478 | 479 | Global Hint Extern 1 (Symmetric (rel_pull _ _ _)) => 480 | eapply rel_pull_sym : typeclass_instances. 481 | 482 | Lemma rel_pull_rcompose {A1 A2 A3 B1 B2 B3} f g h R12 R23 R13 : 483 | @RCompose B1 B2 B3 R12 R23 R13 -> 484 | @RCompose A1 A2 A3 (R12 @@ (f, g)) (R23 @@ (g, h)) (R13 @@ (f, h)). 485 | Proof. 486 | firstorder. 487 | Qed. 488 | 489 | Global Hint Extern 1 (RCompose _ _ (rel_pull _ _ _)) => 490 | eapply rel_pull_rcompose : typeclass_instances. 491 | 492 | (** The introduction rule is straightforward, but changes the head 493 | terms to the functions used with [rel_pull]. This means that an 494 | [RIntro] rule would short-circuit any relational property formulated 495 | for the original terms, even if the codomain is a matching 496 | [rel_pull] relation and we have an appropriate [RElim] instance. 497 | To avoid shadowing such properties, we define our introduction rule 498 | as an [RStep] with a low priority. See test [rel_pull_2]. *) 499 | 500 | Lemma rel_pull_rintro {A B A' B'} (f: A -> A') (g: B -> B') R x y: 501 | RStep (R (f x) (g y)) ((R @@ (f, g))%rel x y). 502 | Proof. 503 | firstorder. 504 | Qed. 505 | 506 | Global Hint Extern 60 (RStep _ ((_ @@ (_, _))%rel _ _)) => 507 | eapply rel_pull_rintro : typeclass_instances. 508 | 509 | (** We can reuse an instance of [RElim] for the underlying relation to 510 | provide corresponding instances for the pulled relation. *) 511 | 512 | Lemma rel_pull_relim {A B A' B'} (f: A -> A') (g: B -> B') R x y P Q: 513 | RElim R (f x) (g y) P Q -> 514 | RElim (R @@ (f, g)) x y P Q. 515 | Proof. 516 | firstorder. 517 | Qed. 518 | 519 | Global Hint Extern 1 (RElim (_ @@ (_, _)) _ _ _ _) => 520 | eapply rel_pull_relim : typeclass_instances. 521 | 522 | (** ** Pushing a relation along functions *) 523 | 524 | Inductive rel_push {A1 A2 B1 B2} f g (R: rel A1 A2): rel B1 B2 := 525 | rel_push_rintro x y: RIntro (R x y) (rel_push f g R) (f x) (g y). 526 | 527 | Global Hint Extern 1 (RIntro _ (rel_push _ _ _) _ _) => 528 | eapply rel_push_rintro : typeclass_instances. 529 | 530 | (** The level we choose here is to match the notation used for 531 | [PTree.get] in Compcert's [Maps] library. *) 532 | 533 | Notation "R !! ( f , g )" := (rel_push f g R) 534 | (at level 1) : rel_scope. 535 | 536 | Notation "R !! f" := (rel_push f f R) 537 | (at level 1) : rel_scope. 538 | 539 | Notation "R !! ( f )" := (rel_push f f R) 540 | (at level 1) : rel_scope. 541 | 542 | Global Instance rel_push_subrel {A1 A2 B1 B2} (f: A1 -> B1) (g: A2 -> B2): 543 | Proper (subrel ++> subrel) (rel_push f g). 544 | Proof. 545 | intros R1 R2 HR x y Hxy. 546 | destruct Hxy. 547 | rintro; eauto. 548 | Qed. 549 | 550 | Global Instance rel_push_subrel_params: 551 | Params (@rel_push) 3 := { }. 552 | 553 | Lemma rel_push_corefl {A B} (f: A -> B) (R: rel A A): 554 | Coreflexive R -> 555 | Coreflexive (R !! f). 556 | Proof. 557 | intros H _ _ [x y Hxy]. 558 | f_equal. 559 | eauto using coreflexivity. 560 | Qed. 561 | 562 | Global Hint Extern 1 (Coreflexive (_ !! _)) => 563 | eapply rel_push_corefl : typeclass_instances. 564 | 565 | (** When using [R !! fst] or [R !! snd], if [rel_push_intro] does not 566 | apply, we can use the following instances instead. *) 567 | 568 | Lemma rel_push_fst_rexists {A1 A2 B1 B2} (x1:A1) (x2:A2) (y1:B1) (y2:B2) R: 569 | RExists (R (x1, y1) (x2, y2)) (R !! fst) x1 x2. 570 | Proof. 571 | intros H. 572 | change x1 with (fst (x1, y1)). 573 | change x2 with (fst (x2, y2)). 574 | rintro. 575 | assumption. 576 | Qed. 577 | 578 | Global Hint Extern 1 (RExists _ (_ !! fst) _ _) => 579 | eapply rel_push_fst_rexists : typeclass_instances. 580 | 581 | Lemma rel_push_snd_rexists {A1 A2 B1 B2} (x1:A1) (x2:A2) (y1:B1) (y2:B2) R: 582 | RExists (R (x1, y1) (x2, y2)) (R !! snd) y1 y2. 583 | Proof. 584 | intros H. 585 | change y1 with (snd (x1, y1)). 586 | change y2 with (snd (x2, y2)). 587 | rintro. 588 | assumption. 589 | Qed. 590 | 591 | Global Hint Extern 1 (RExists _ (_ !! snd) _ _) => 592 | eapply rel_push_snd_rexists : typeclass_instances. 593 | 594 | (** ** Relation currying *) 595 | 596 | (* An interesting special case of [rel_pull] is [rel_curry], which 597 | takes a relation over [A * B -> C] and turns it into a relation 598 | over [A -> B -> C]. 599 | 600 | [rel_curry] is particularly useful when two (curried) arguments 601 | to a given function have to be related in a dependent way. For 602 | example in Compcert, memory injections relate memory blocks and 603 | offsets jointly, but many operations have those as two separate 604 | arguments. To state the relational property of such operations, we 605 | can uncurry a joint relation on (block, offset) pairs. 606 | 607 | [rel_curry] can be obtained by pulling back the original relation 608 | along [uncurry]: *) 609 | 610 | Definition curry {A B C} (f: A * B -> C): A -> B -> C := 611 | fun a b => f (a, b). 612 | 613 | Definition uncurry {A B C} (f: A -> B -> C): A * B -> C := 614 | fun ab => match ab with (a, b) => f a b end. 615 | 616 | Definition rel_curry {A1 B1 C1 A2 B2 C2} (R: rel (A1*B1->C1) (A2*B2->C2)) := 617 | (R @@ uncurry)%rel. 618 | 619 | Definition rel_uncurry {A1 B1 C1 A2 B2 C2} (R: rel (A1->B1->C1) (A2->B2->C2)) := 620 | (R @@ curry)%rel. 621 | 622 | (** We use the following notation for [rel_curry], which evokes 623 | splitting a pair into two separate arguments. Note that we use the 624 | same priority as for [++>], [-->], and [==>], because we usually 625 | want [rel_curry] to nest in such a way that for instance, 626 | [S ++> % R --> % R ++> impl] will be interpreted as 627 | [S ++> rel_curry (R --> rel_curry (R ++> impl))]. *) 628 | 629 | Notation "% R" := (rel_curry R) (at level 55, right associativity) : rel_scope. 630 | 631 | (** In order to provide an [RElim] instance for [rel_curry], we will 632 | rely on the fact that: 633 | 634 | uncurry f (x1, x2) x3 ... xn = f x1 x2 x3 ... xn, 635 | 636 | and that: 637 | 638 | rel_curry R f g <-> R (uncurry f) (uncurry g), 639 | 640 | so that when the instance of [RElim] associated with [R] allows us 641 | to show: 642 | 643 | Rcod (uncurry f (x1, x2) x3 ... xn) (uncurry g (y1, y2) y3 ... yn) 644 | 645 | we can convert it into an instance for proving: 646 | 647 | Rcod (f x1 x2 x3 ... xn) (g y1 y2 y3 ... yn), 648 | 649 | which is what will be expected for [monotonicity]. To accomplish 650 | this, we need to [unfold uncurry] in the result provided for [R] 651 | before we use it to solve the [rel_uncurry] case. This helper class 652 | does this. *) 653 | 654 | Class UnfoldUncurry {A} (before: A) (after: A) := 655 | unfold_uncurry_before_after: before = after. 656 | 657 | Ltac unfold_uncurry := 658 | match goal with 659 | | |- context C[uncurry ?f ?p] => 660 | is_evar p; 661 | let T := type of p in 662 | let Av := fresh in evar (Av: Type); 663 | let A := eval red in Av in clear Av; 664 | let Bv := fresh in evar (Bv: Type); 665 | let B := eval red in Bv in clear Bv; 666 | unify T (prod A B)%type; 667 | let av := fresh in evar (av : A); 668 | let a := eval red in av in clear av; 669 | let bv := fresh in evar (bv : B); 670 | let b := eval red in bv in clear bv; 671 | let G := context C[f a b] in 672 | unify p (a, b); 673 | change G 674 | end. 675 | 676 | Global Hint Extern 1 (UnfoldUncurry ?P ?Q) => 677 | repeat unfold_uncurry; constructor : typeclass_instances. 678 | 679 | (** Now we can provide a somewhat general [RElim] instance for 680 | [rel_uncurry]. *) 681 | 682 | Lemma rel_curry_relim {A1 B1 C1 A2 B2 C2} R f g P Q Q': 683 | @RElim (A1 * B1 -> C1) (A2 * B2 -> C2) R (uncurry f) (uncurry g) P Q -> 684 | UnfoldUncurry Q Q' -> 685 | @RElim (A1 -> B1 -> C1) (A2 -> B2 -> C2) (% R) f g P Q'. 686 | Proof. 687 | unfold UnfoldUncurry. 688 | intros; subst. 689 | assumption. 690 | Qed. 691 | 692 | Global Hint Extern 60 (RStep _ ((% _)%rel _ _)) => 693 | eapply rel_pull_rintro : typeclass_instances. 694 | 695 | Global Hint Extern 1 (RElim (% _) _ _ _ _) => 696 | eapply rel_curry_relim : typeclass_instances. 697 | 698 | (** ** The [req] relation *) 699 | 700 | (** The relation [req a] asserts that both elements are equal to [a]. 701 | This comes in handy when expressing the relational properties of 702 | functions: it can be used to fix the value of an argument, or 703 | parametrize it using a variable with a broader scope. *) 704 | 705 | Inductive req {A} (a: A): rel A A := 706 | req_intro: req a a a. 707 | 708 | Lemma req_rintro {A} (a: A): 709 | RIntro True (req a) a a. 710 | Proof. 711 | firstorder. 712 | Qed. 713 | 714 | Global Hint Extern 0 (RIntro _ (req _) _ _) => 715 | eapply req_rintro : typeclass_instances. 716 | 717 | Lemma req_corefl {A} (a: A): 718 | Coreflexive (req a). 719 | Proof. 720 | destruct 1. 721 | reflexivity. 722 | Qed. 723 | 724 | Global Hint Extern 0 (Coreflexive (req _)) => 725 | eapply req_corefl : typeclass_instances. 726 | 727 | (** ** Checking predicates on the left and right elements *) 728 | 729 | Definition lsat {A B} (P: A -> Prop): rel A B := 730 | fun x y => P x. 731 | 732 | Global Instance lsat_subrel A B: 733 | Monotonic (@lsat A B) ((- ==> impl) ++> subrel). 734 | Proof. 735 | firstorder. 736 | Qed. 737 | 738 | Global Instance lsat_subrel_params: 739 | Params (@lsat) 3 := { }. 740 | 741 | Definition rsat {A B} (P: B -> Prop): rel A B := 742 | fun x y => P y. 743 | 744 | Global Instance rsat_subrel A B: 745 | Monotonic (@rsat A B) ((- ==> impl) ++> subrel). 746 | Proof. 747 | firstorder. 748 | Qed. 749 | 750 | Global Instance rsat_subrel_params: 751 | Params (@rsat) 3 := { }. 752 | 753 | Inductive psat {A} (I: A -> Prop) (x: A): A -> Prop := 754 | psat_intro: I x -> psat I x x. 755 | 756 | Global Instance psat_subrel A: 757 | Monotonic (@psat A) ((- ==> impl) ++> subrel). 758 | Proof. 759 | intros P Q HPQ x _ [Hx]. 760 | constructor. apply HPQ. assumption. 761 | Qed. 762 | 763 | Global Instance psat_subrel_params: 764 | Params (@psat) 3 := { }. 765 | 766 | Lemma psat_corefl {A} (I: A -> Prop): 767 | Coreflexive (psat I). 768 | Proof. 769 | intros x _ [_]. reflexivity. 770 | Qed. 771 | 772 | Global Hint Extern 0 (Coreflexive (psat _)) => 773 | eapply psat_corefl : typeclass_instances. 774 | 775 | (** ** Relation versions of [ex] and [all] *) 776 | 777 | (** Ideally we would overload the [forall] and [exists] notation to 778 | use the relation version under the [rel] scope. But as long as 779 | we keep [rel_scope] open globally, we can't really do that 780 | without breaking everything. So we use our own keyword [rexists] and 781 | [rforall] instead. *) 782 | 783 | Definition rel_all {A B C} (R: C -> rel A B): rel A B := 784 | fun x y => forall c, R c x y. 785 | 786 | Notation "'rforall' x .. y , p" := 787 | (rel_all (fun x => .. (rel_all (fun y => p%rel)) .. )) 788 | (at level 200, x binder, right associativity) 789 | : rel_scope. 790 | 791 | Lemma rel_all_rintro {A B C} (R: C -> rel A B) m n: 792 | RIntro (forall c : C, R c m n) (rel_all R) m n. 793 | Proof. 794 | firstorder. 795 | Qed. 796 | 797 | Global Hint Extern 0 (RIntro _ (rel_all _) _ _) => 798 | eapply rel_all_rintro : typeclass_instances. 799 | 800 | Lemma rel_all_relim {A B C} (R: C -> rel A B) x y P Q: 801 | (exists c, RElim (R c) x y P Q) -> 802 | RElim (rel_all R) x y P Q. 803 | Proof. 804 | firstorder. 805 | Qed. 806 | 807 | Global Hint Extern 1 (RElim (rel_all _) _ _ _ _) => 808 | eapply rel_all_relim; eexists : typeclass_instances. 809 | 810 | Definition rel_ex {A B C} (R: C -> rel A B): rel A B := 811 | fun x y => exists c, R c x y. 812 | 813 | Notation "'rexists' x .. y , p" := 814 | (rel_ex (fun x => .. (rel_ex (fun y => p%rel)) ..)) 815 | (at level 200, x binder, right associativity) 816 | : rel_scope. 817 | 818 | Lemma rel_ex_rintro {A B C} (R: C -> rel A B) c m n: 819 | RExists (R c m n) (rel_ex R) m n. 820 | Proof. 821 | firstorder. 822 | Qed. 823 | 824 | Global Hint Extern 0 (RExists _ (rel_ex _) _ _) => 825 | eapply rel_ex_rintro : typeclass_instances. 826 | 827 | Lemma rel_ex_relim {A B C} (R: C -> rel A B) x y P Q: 828 | (forall c, RElim (R c) x y P Q) -> 829 | RElim (rel_ex R) x y P Q. 830 | Proof. 831 | firstorder. 832 | Qed. 833 | 834 | Global Hint Extern 1 (RElim (rel_ex _) _ _ _ _) => 835 | eapply rel_ex_relim : typeclass_instances. 836 | 837 | (** ** The [rel_incr] construction *) 838 | 839 | (** XXX: this is on the way out, see KLR.v *) 840 | 841 | (** When dealing with Kripke logical relations, we have a family of 842 | relations indexed by a type of worlds, as well as an accessibility 843 | relation over this type of worlds. We often want to state that there 844 | exists a world accessible from a base one in which the relation 845 | holds. The following construction expresses this. *) 846 | 847 | Definition rel_incr {W A B} (acc: rel W W) (R: W -> rel A B): W -> rel A B := 848 | fun w a b => exists w', acc w w' /\ R w' a b. 849 | 850 | (** While it's possible to define a more general monotonicity property 851 | for [rel_incr], this one is well-behaved and usually corresponds to 852 | what we want. *) 853 | 854 | Global Instance rel_incr_subrel {W A B} acc: 855 | Transitive acc -> 856 | Monotonic (@rel_incr W A B acc) ((- ==> subrel) ++> acc --> subrel). 857 | Proof. 858 | intros Hacc R1 R2 HR w1 w2 Hw a b (w1' & Hw1' & Hab). 859 | eexists; split. 860 | - transitivity w1; 861 | eassumption. 862 | - apply HR. 863 | assumption. 864 | Qed. 865 | 866 | Global Instance rel_incr_subrel_params: 867 | Params (@rel_incr) 4 := { }. 868 | 869 | (** Note the order of the premises in our intro rule. We want to first 870 | determine what [w'] should be, then prove [acc w w']. *) 871 | 872 | Lemma rel_incr_rintro {W A B} (acc: rel W W) (R: W -> rel A B) w w' m n: 873 | RExists (R w' m n /\ acc w w') (rel_incr acc R w) m n. 874 | Proof. 875 | firstorder. 876 | Qed. 877 | 878 | Global Hint Extern 0 (RExists _ (rel_incr _ _ _) _ _) => 879 | eapply rel_incr_rintro : typeclass_instances. 880 | 881 | Lemma rel_incr_rdestruct {W A B} acc R w T: 882 | (forall w, exists Tw, RDestruct (R w) Tw /\ Convertible (T w) Tw) -> 883 | RDestruct 884 | (@rel_incr W A B acc R w) 885 | (fun P => forall w', acc w w' -> Delay.unpack (T w' P)). 886 | Proof. 887 | clear. 888 | intros HR m n (w' & Hw' & Hmn) P H. 889 | destruct (HR w') as (Tw & HRw' & HTw). 890 | eapply rdestruct; eauto. 891 | destruct HTw. 892 | eapply H; eauto. 893 | Qed. 894 | 895 | Global Hint Extern 2 (RDestruct (rel_incr _ _ _) _) => 896 | eapply rel_incr_rdestruct; intro; eexists; split : typeclass_instances. 897 | -------------------------------------------------------------------------------- /Relators.v: -------------------------------------------------------------------------------- 1 | Require Export RelDefinitions. 2 | Require Import RelClasses. 3 | Require Import Coq.Lists.List. 4 | 5 | 6 | (** * Relators *) 7 | 8 | (** With this infrastructure in place, we can define actual relators 9 | that cover the commonly used type constructors. There are two broad 10 | categories: those related to function types, and those derived from 11 | inductive types. *) 12 | 13 | (** As a convention, we name relators and the associated monotonicity 14 | theorems by appending the suffix [_rel] to the name of original type 15 | and those of its constructors. Likewise, we use the suffix [_subrel] 16 | for monotonicity theorems that characterize the variance of 17 | corresponding relators with respect to [subrel], and the suffix 18 | [_relim] for an associated instance of [RElim]. *) 19 | 20 | (** ** A note about universes *) 21 | 22 | (** Our use of [subrel] illustrates the fact that we consider 23 | relations over relations, and use relators to build up relations 24 | over the types of our relators to characterize their variance. 25 | In other words: we relate all kind of higher-order things as a 26 | matter of course and make sure our library eats its own dog 27 | food. This makes it relatively easy to run into universe 28 | inconsistency issues. This note is to explain why those happen, fix 29 | the terminology, and formulate best practices to avoid such 30 | problems. 31 | 32 | The critical universe level in Coqrel is the universe [U] from which 33 | the arguments and return type of [rel] are taken (so that 34 | conceptually we will write [rel : U -> U -> U]). We will call 35 | "small types" the inhabitants of [U], and "large types" those taken 36 | from a universe strictly larger than [U]. At the moment, [U] is 37 | implicit, but at some point we may want declare an explicit 38 | [Universe] to aid debugging and enhance clarity. 39 | 40 | Generally speaking, [U] is a pretty large universe: for the most 41 | part, the world of relational things sits well above the world of 42 | the things we want relate, no matter how high-order those are. 43 | Even [rel A B] is still a small type, so that 44 | [@subrel : forall A B : U, rel (rel A B) (rel A B)] 45 | above didn't cause any issue, although the type of [rel] itself, 46 | namely [U -> U -> U : U+1], is of course a large type. 47 | 48 | With this in mind, consider the relational operator 49 | [rel_compose : forall A B C, rel A B -> rel B C -> rel A C]. 50 | We wish to register a monotonicity property for [rel_compose], 51 | establishing that composing larger relations yields a larger 52 | composite relation. As usual, we express this property by 53 | establishing that [rel_compose] is related to itself by an 54 | appropriate relation. However, the type of [rel_compose] may involve 55 | [U : U+1] as the type of the arguments [A B C], so that the type of 56 | [rel_compose] may itself be forced into sort [U+1]. Consequently, 57 | attempting to define that relation will yield a universe 58 | inconsistency as soon as, say, [subrel] is used as an argument of 59 | [rel_compose]. 60 | 61 | Fortunately, even if [A B C : U], the specialized form 62 | [rel_compose A B C] is still in the small type 63 | [rel A B -> rel B C -> rel A C : U]. Moreover, there is no 64 | relational structure that we want to capture regarding the type 65 | parameters [A B C], so that we don't lose anything by defining our 66 | monotonicity property for the specialized [rel_compose A B C], 67 | as long as the property itself is generic in the specific values of 68 | [A B C]. 69 | 70 | This is why, as a general rule, we define the [foo_subrel] 71 | properties of our relators on partially applied versions thereof, 72 | and use a [foo_subrel_params] hint to make sure our monotonicity 73 | tactic can resolve them. *) 74 | 75 | (** ** Relators for function types *) 76 | 77 | (** *** Pointwise extension of a relation *) 78 | 79 | (** One useful special case is the pointwise extension of a relation 80 | on the domain to the function type. This is equivalent to [eq ==> R], 81 | however with the formulation below we don't have consider two equal 82 | elements of the domain. *) 83 | 84 | Definition arrow_pointwise_rel A {B1 B2}: 85 | rel B1 B2 -> rel (A -> B1) (A -> B2) := 86 | fun RB f g => forall a, RB (f a) (g a). 87 | 88 | Arguments arrow_pointwise_rel A%type {B1%type B2%type} RB%rel _ _. 89 | 90 | Notation "- ==> R" := (arrow_pointwise_rel _ R) 91 | (at level 55, right associativity) : rel_scope. 92 | 93 | Global Instance arrow_pointwise_subrel {A B1 B2}: 94 | Monotonic (@arrow_pointwise_rel A B1 B2) (subrel ++> subrel). 95 | Proof. 96 | firstorder. 97 | Qed. 98 | 99 | Global Instance arrow_pointwise_subrel_params: 100 | Params (@arrow_pointwise_rel) 3 := { }. 101 | 102 | Lemma arrow_pointwise_rintro {A B1 B2} (R: rel B1 B2) f g: 103 | RIntro (forall x: A, R (f x) (g x)) (- ==> R) f g. 104 | Proof. 105 | firstorder. 106 | Qed. 107 | 108 | Global Hint Extern 0 (RIntro _ (- ==> _) _ _) => 109 | eapply arrow_pointwise_rintro : typeclass_instances. 110 | 111 | (** Note that although the elimination rule could use a single 112 | variable and dispense with the equality premise, it actually uses 113 | two distinct variables [m] and [n], and a premise in [P] that [m = n]. 114 | Hence the effect is similar to the elimination rule for the 115 | equivalent relation [eq ==> R]. We do this because in contexts where 116 | we have a single term [t] against which both [m] and [n] unify, it 117 | is easy and cheap to automatically prove that [t = t], but in 118 | contexts where we have two disctinct variables and a proof of 119 | equality, it is more involved to massage the goal into a form where 120 | the single-variable version of the rule could apply. *) 121 | 122 | Lemma arrow_pointwise_relim {A B1 B2} (R: rel B1 B2) f g (m n: A) P Q: 123 | RElim R (f m) (g n) P Q -> 124 | RElim (- ==> R) f g (m = n /\ P) Q. 125 | Proof. 126 | intros HPQ Hfg [Hab HP]. 127 | subst. 128 | firstorder. 129 | Qed. 130 | 131 | Global Hint Extern 1 (RElim (- ==> _) _ _ _ _) => 132 | eapply arrow_pointwise_relim : typeclass_instances. 133 | 134 | Lemma arrow_pointwise_refl {T} `(Reflexive) : 135 | @Reflexive (T -> A) (- ==> R). 136 | Proof. 137 | firstorder. 138 | Qed. 139 | 140 | Global Hint Extern 1 (Reflexive (- ==> _)) => 141 | eapply arrow_pointwise_refl : typeclass_instances. 142 | 143 | Global Instance arrow_pointwise_rel_compose {T} `(RCompose) : 144 | RCompose (A := T -> A) (- ==> RAB) (- ==> RBC) (- ==> RAC). 145 | Proof. 146 | firstorder. 147 | Qed. 148 | 149 | (** *** Dependent pointwise extension *) 150 | 151 | (** Like we did for non-dependent functions, we can provide a simpler 152 | definition for the special case where [E] is [eq]. *) 153 | 154 | Definition forall_pointwise_rel {V: Type} {FV1 FV2: V -> Type}: 155 | (forall v, rel (FV1 v) (FV2 v)) -> 156 | rel (forall v, FV1 v) (forall v, FV2 v) := 157 | fun FE f g => 158 | forall v, FE v (f v) (g v). 159 | 160 | Arguments forall_pointwise_rel {V%type FV1%type FV2%type} FE%rel _ _. 161 | 162 | Notation "'forallr' - @ v : V , FE" := 163 | (forall_pointwise_rel (V := V) (fun v => FE)) 164 | (v ident, at level 200). 165 | 166 | Notation "'forallr' - @ v , FE" := 167 | (forall_pointwise_rel (fun v => FE)) 168 | (v ident, at level 200). 169 | 170 | Notation "'forallr' - : V , FE" := 171 | (forall_pointwise_rel (V := V) (fun _ => FE)) 172 | (at level 200). 173 | 174 | Notation "'forallr' - , FE" := 175 | (forall_pointwise_rel (fun _ => FE)) 176 | (at level 200). 177 | 178 | Lemma forall_pointwise_rintro {V FV1 FV2} (FE: forall v, rel _ _) f g: 179 | RIntro 180 | (forall v, FE v (f v) (g v)) 181 | (@forall_pointwise_rel V FV1 FV2 FE) f g. 182 | Proof. 183 | firstorder. 184 | Qed. 185 | 186 | Global Hint Extern 0 (RIntro _ (forall_pointwise_rel _) _ _) => 187 | eapply forall_pointwise_rintro : typeclass_instances. 188 | 189 | Lemma forall_pointwise_relim {V FV1 FV2} R f g v P Q: 190 | RElim (R v) (f v) (g v) P Q -> 191 | RElim (@forall_pointwise_rel V FV1 FV2 R) f g P Q. 192 | Proof. 193 | firstorder. 194 | Qed. 195 | 196 | Global Hint Extern 1 (RElim (forall_pointwise_rel _) _ _ _ _) => 197 | eapply forall_pointwise_relim : typeclass_instances. 198 | 199 | (** *** Dependent products (restricted version) *) 200 | 201 | (** This is a restricted version of [forall_rel] with [E] in [Prop], 202 | when the codomain relation only depends on the arguments. *) 203 | 204 | Definition forallp_rel {V1 V2} (E: rel V1 V2) {FV1: V1->Type} {FV2: V2->Type}: 205 | (forall v1 v2, rel (FV1 v1) (FV2 v2)) -> 206 | rel (forall v1, FV1 v1) (forall v2, FV2 v2) := 207 | fun FE f g => 208 | forall v1 v2, E v1 v2 -> FE v1 v2 (f v1) (g v2). 209 | 210 | Arguments forallp_rel {V1%type V2%type} E%rel {FV1%type FV2%type} FE%rel _ _. 211 | 212 | Notation "'forallr' v1 v2 : E , R" := 213 | (forallp_rel E (fun v1 v2 => R)) 214 | (at level 200, v1 ident, v2 ident, right associativity) 215 | : rel_scope. 216 | 217 | Lemma forallp_rintro {V1 V2} {E: rel V1 V2} {F1 F2} (FE: forall v1 v2, rel _ _) f g: 218 | RIntro 219 | (forall v1 v2, E v1 v2 -> FE v1 v2 (f v1) (g v2)) 220 | (@forallp_rel V1 V2 E F1 F2 FE) f g. 221 | Proof. 222 | firstorder. 223 | Qed. 224 | 225 | Global Hint Extern 0 (RIntro _ (forallp_rel _ _) _ _) => 226 | eapply forallp_rintro : typeclass_instances. 227 | 228 | (** Since [e : E v1 v2] cannot be unified in [Q], the elimination rule 229 | adds an [E v1 v2] premise to [P] instead. *) 230 | 231 | Lemma forallp_relim {V1 V2 E FV1 FV2} R f g v1 v2 P Q: 232 | RElim (R v1 v2) (f v1) (g v2) P Q -> 233 | RElim (@forallp_rel V1 V2 E FV1 FV2 R) f g (E v1 v2 /\ P) Q. 234 | Proof. 235 | firstorder. 236 | Qed. 237 | 238 | Global Hint Extern 1 (RElim (forallp_rel _ _) _ _ _ _) => 239 | eapply forallp_relim : typeclass_instances. 240 | 241 | (** TODO: A further specialization could be [foralln_rel], which does 242 | not even require that [v1], [v2] be related (let alone care about 243 | the proof). In other words, a [forallr v1 v2, R] which would 244 | essentially boil down to [forallr v1 v2 : ⊤, R]. *) 245 | 246 | (** *** Sets ([A -> Prop]) *) 247 | 248 | (** Sets of type [A -> Prop] can be related using the regular arrow 249 | relator, as in [R ++> impl]. This states that for any two elements 250 | related by [R], if the first is in the first set, then the second 251 | must be in the second set. 252 | 253 | However, very often we need the following relator instead, which 254 | states that for any element of the first set, there exists an 255 | element of the second set which is related to it. This is useful for 256 | example when defining simulation diagrams. *) 257 | 258 | Definition set_le {A B} (R: rel A B): rel (A -> Prop) (B -> Prop) := 259 | fun sA sB => forall a, sA a -> exists b, sB b /\ R a b. 260 | 261 | Global Instance set_le_subrel {A B}: 262 | Monotonic (@set_le A B) (subrel ++> subrel). 263 | Proof. 264 | intros R1 R2 HR sA sB Hs. 265 | intros x Hx. 266 | destruct (Hs x) as (y & Hy & Hxy); eauto. 267 | Qed. 268 | 269 | Global Instance set_le_subrel_params: 270 | Params (@set_le) 3 := { }. 271 | 272 | Lemma set_le_refl {A} (R : relation A) : 273 | Reflexive R -> 274 | Reflexive (set_le R). 275 | Proof. 276 | firstorder. 277 | Qed. 278 | 279 | Global Hint Extern 1 (Reflexive (set_le _)) => 280 | eapply set_le_refl : typeclass_instances. 281 | 282 | Global Instance set_le_compose `(RCompose) : 283 | RCompose (set_le RAB) (set_le RBC) (set_le RAC). 284 | Proof. 285 | intros sa sb sc Hsab Hsbc a Ha. 286 | edestruct Hsab as (b & Hb & Hab); eauto. 287 | edestruct Hsbc as (c & Hc & Hbc); eauto. 288 | Qed. 289 | 290 | (** We define [set_ge] as well. *) 291 | 292 | Definition set_ge {A B} (R: rel A B): rel (A -> Prop) (B -> Prop) := 293 | fun sA sB => forall b, sB b -> exists a, sA a /\ R a b. 294 | 295 | Global Instance set_ge_subrel {A B}: 296 | Monotonic (@set_ge A B) (subrel ++> subrel). 297 | Proof. 298 | firstorder. 299 | Qed. 300 | 301 | Global Instance set_ge_subrel_params: 302 | Params (@set_ge) 3 := { }. 303 | 304 | Lemma set_ge_refl {A} (R : relation A) : 305 | Reflexive R -> 306 | Reflexive (set_ge R). 307 | Proof. 308 | firstorder. 309 | Qed. 310 | 311 | Global Hint Extern 1 (Reflexive (set_ge _)) => 312 | eapply set_ge_refl : typeclass_instances. 313 | 314 | Global Instance set_ge_compose `(RCompose) : 315 | RCompose (set_ge RAB) (set_ge RBC) (set_ge RAC). 316 | Proof. 317 | intros sa sb sc Hsab Hsbc c Hc. 318 | edestruct Hsbc as (b & Hb & Hbc); eauto. 319 | edestruct Hsab as (a & Ha & Hab); eauto. 320 | Qed. 321 | 322 | (** ** Inductive types *) 323 | 324 | (** For inductive types, there is a systematic way of converting their 325 | definition into that of the corresponding relator. Where the 326 | original inductive definition quantifies over types, the 327 | corresponding relator will quantify over pairs of types and 328 | relations between them. Then, the constructors of the relator will 329 | essentially be [Proper] instances for the original constructors. 330 | In other words, the resulting relation will be the smallest one such 331 | that the constructors are order-preserving. *) 332 | 333 | (** *** Nullary type constructors *) 334 | 335 | (** As a proof-of-concept, we start with the most elementary types 336 | [Empty_set] and [unit], which can be considered as nullary type 337 | constructors related to [sum] and [prod] below. *) 338 | 339 | Inductive Empty_set_rel: rel Empty_set Empty_set := . 340 | 341 | Inductive unit_rel: rel unit unit := 342 | tt_rel: Proper unit_rel tt. 343 | 344 | Global Existing Instance tt_rel. 345 | 346 | (** *** Sum types *) 347 | 348 | (** The definition of [sum_rel] could look something like this: 349 | << 350 | Inductive sum_rel: 351 | forall {A1 A2 B1 B2}, rel A1 A2 -> rel B1 B2 -> rel (A1+B1) (A2+B2):= 352 | | inl_rel: Proper (∀ RA : rel, ∀ RB : rel, RA ++> sum_rel RA RB) (@inl) 353 | | inr_rel: Proper (∀ RA : rel, ∀ RB : rel, RB ++> sum_rel RA RB) (@inr). 354 | >> 355 | However, to minimize the need for [inversion]s we want to keep as 356 | many arguments as possible as parameters of the inductive type. *) 357 | 358 | Inductive sum_rel {A1 A2 B1 B2} RA RB: rel (A1 + B1)%type (A2 + B2)%type := 359 | | inl_rel_def: (RA ++> sum_rel RA RB)%rel (@inl A1 B1) (@inl A2 B2) 360 | | inr_rel_def: (RB ++> sum_rel RA RB)%rel (@inr A1 B1) (@inr A2 B2). 361 | 362 | Infix "+" := sum_rel : rel_scope. 363 | 364 | (** Since it is not possible to retype the constructors after the 365 | fact, we use the [_def] suffix when defining them, then redeclare 366 | a corresponding, full-blown [Proper] instance. *) 367 | 368 | Global Instance inl_rel: 369 | Monotonic (@inl) (forallr RA, forallr RB, RA ++> RA + RB). 370 | Proof. 371 | repeat intro; apply inl_rel_def; assumption. 372 | Qed. 373 | 374 | Global Instance inr_rel: 375 | Monotonic (@inr) (forallr RA, forallr RB, RB ++> RA + RB). 376 | Proof. 377 | repeat intro; apply inr_rel_def; assumption. 378 | Qed. 379 | 380 | Global Instance sum_subrel {A1 A2 B1 B2}: 381 | Monotonic (@sum_rel A1 A2 B1 B2) (subrel ++> subrel ++> subrel). 382 | Proof. 383 | intros RA1 RA2 HRA RB1 RB2 HRB. 384 | intros x1 x2 Hx. 385 | destruct Hx; constructor; eauto. 386 | Qed. 387 | 388 | Global Instance sum_subrel_params: 389 | Params (@sum_rel) 4 := { }. 390 | 391 | Lemma sum_rel_refl {A B} (R1: rel A A) (R2: rel B B): 392 | Reflexive R1 -> Reflexive R2 -> Reflexive (R1 + R2). 393 | Proof. 394 | intros H1 H2 x. 395 | destruct x; constructor; reflexivity. 396 | Qed. 397 | 398 | Global Hint Extern 2 (Reflexive (_ + _)) => 399 | eapply sum_rel_refl : typeclass_instances. 400 | 401 | Lemma sum_rel_corefl {A B} (R1: rel A A) (R2: rel B B): 402 | Coreflexive R1 -> Coreflexive R2 -> Coreflexive (R1 + R2). 403 | Proof. 404 | intros H1 H2 _ _ [x y H | x y H]; 405 | f_equal; 406 | eauto using coreflexivity. 407 | Qed. 408 | 409 | Global Hint Extern 2 (Coreflexive (_ + _)) => 410 | eapply sum_rel_corefl : typeclass_instances. 411 | 412 | Lemma sum_rel_trans {A B} (R1: rel A A) (R2: rel B B): 413 | Transitive R1 -> Transitive R2 -> Transitive (R1 + R2). 414 | Proof. 415 | intros H1 H2 x y z Hxy Hyz. 416 | destruct Hxy; inversion Hyz; constructor; etransitivity; eassumption. 417 | Qed. 418 | 419 | Global Hint Extern 2 (Transitive (_ + _)) => 420 | eapply sum_rel_trans : typeclass_instances. 421 | 422 | Lemma sum_rel_sym {A B} (R1: rel A A) (R2: rel B B): 423 | Symmetric R1 -> Symmetric R2 -> Symmetric (R1 + R2). 424 | Proof. 425 | intros H1 H2 x y Hxy. 426 | destruct Hxy; constructor; symmetry; eassumption. 427 | Qed. 428 | 429 | Global Hint Extern 2 (Symmetric (_ + _)) => 430 | eapply sum_rel_sym : typeclass_instances. 431 | 432 | Lemma sum_rel_preorder {A B} (R1: rel A A) (R2: rel B B): 433 | PreOrder R1 -> PreOrder R2 -> PreOrder (R1 + R2). 434 | Proof. 435 | split; typeclasses eauto. 436 | Qed. 437 | 438 | Global Hint Extern 2 (PreOrder (_ + _)) => 439 | eapply sum_rel_preorder : typeclass_instances. 440 | 441 | (** *** Pairs *) 442 | 443 | Definition prod_rel {A1 A2 B1 B2} RA RB: rel (A1 * B1)%type (A2 * B2)%type := 444 | fun x1 x2 => RA (fst x1) (fst x2) /\ RB (snd x1) (snd x2). 445 | 446 | Infix "*" := prod_rel : rel_scope. 447 | 448 | Global Instance pair_rel: 449 | Monotonic (@pair) (forallr RA, forallr RB, RA ++> RB ++> RA * RB). 450 | Proof. 451 | intros A1 A2 RA B1 B2 RB a1 a2 Ha b1 b2 Hb. 452 | red. 453 | eauto. 454 | Qed. 455 | 456 | Global Instance fst_rel: 457 | Monotonic (@fst) (forallr RA, forallr RB, RA * RB ==> RA). 458 | Proof. 459 | intros A1 A2 RA B1 B2 RB. 460 | intros x1 x2 [Ha Hb]. 461 | assumption. 462 | Qed. 463 | 464 | Global Instance snd_rel: 465 | Monotonic (@snd) (forallr RA, forallr RB, RA * RB ==> RB). 466 | Proof. 467 | intros A1 A2 RA B1 B2 RB. 468 | intros x1 x2 [Ha Hb]. 469 | assumption. 470 | Qed. 471 | 472 | Global Instance prod_subrel {A1 A2 B1 B2}: 473 | Monotonic (@prod_rel A1 A2 B1 B2) (subrel ++> subrel ++> subrel). 474 | Proof. 475 | firstorder. 476 | Qed. 477 | 478 | Global Instance prod_subrel_params: 479 | Params (@prod_rel) 4 := { }. 480 | 481 | Global Instance prod_rdestruct {A1 B1 A2 B2} (RA: rel A1 A2) (RB: rel B1 B2): 482 | RDestruct 483 | (RA * RB)%rel 484 | (fun P => forall a1 a2 b1 b2, RA a1 a2 -> RB b1 b2 -> P (a1, b1) (a2, b2)). 485 | Proof. 486 | intros [a1 b1] [a2 b2] [Ha Hb] P HP. 487 | firstorder. 488 | Qed. 489 | 490 | Lemma prod_rel_refl {A B} (R1: rel A A) (R2: rel B B): 491 | Reflexive R1 -> Reflexive R2 -> Reflexive (R1 * R2). 492 | Proof. 493 | intros H1 H2 x. 494 | destruct x; constructor; reflexivity. 495 | Qed. 496 | 497 | Global Hint Extern 2 (Reflexive (_ * _)) => 498 | eapply prod_rel_refl : typeclass_instances. 499 | 500 | Lemma prod_rel_corefl {A B} (R1: rel A A) (R2: rel B B): 501 | Coreflexive R1 -> Coreflexive R2 -> Coreflexive (R1 * R2). 502 | Proof. 503 | intros H1 H2 [a1 b1] [a2 b2] [Ha Hb]. 504 | f_equal; eauto using coreflexivity. 505 | Qed. 506 | 507 | Global Hint Extern 2 (Coreflexive (_ * _)) => 508 | eapply prod_rel_corefl : typeclass_instances. 509 | 510 | Lemma prod_rel_trans {A B} (R1: rel A A) (R2: rel B B): 511 | Transitive R1 -> Transitive R2 -> Transitive (R1 * R2). 512 | Proof. 513 | intros H1 H2 x y z Hxy Hyz. 514 | destruct Hxy; inversion Hyz; constructor; etransitivity; eassumption. 515 | Qed. 516 | 517 | Global Hint Extern 2 (Transitive (_ * _)) => 518 | eapply prod_rel_trans : typeclass_instances. 519 | 520 | Lemma prod_rel_sym {A B} (R1: rel A A) (R2: rel B B): 521 | Symmetric R1 -> Symmetric R2 -> Symmetric (R1 * R2). 522 | Proof. 523 | intros H1 H2 x y Hxy. 524 | destruct Hxy; constructor; symmetry; eassumption. 525 | Qed. 526 | 527 | Global Hint Extern 2 (Symmetric (_ * _)) => 528 | eapply prod_rel_sym : typeclass_instances. 529 | 530 | Lemma prod_rel_preorder {A B} (R1: rel A A) (R2: rel B B): 531 | PreOrder R1 -> PreOrder R2 -> PreOrder (R1 * R2). 532 | Proof. 533 | split; typeclasses eauto. 534 | Qed. 535 | 536 | Global Hint Extern 2 (PreOrder (_ * _)) => 537 | eapply prod_rel_preorder : typeclass_instances. 538 | 539 | (** *** Option types *) 540 | 541 | Inductive option_rel {A1 A2} (RA: rel A1 A2): rel (option A1) (option A2) := 542 | | Some_rel_def: (RA ++> option_rel RA)%rel (@Some A1) (@Some A2) 543 | | None_rel_def: option_rel RA (@None A1) (@None A2). 544 | 545 | Global Instance Some_rel: 546 | Monotonic (@Some) (forallr R @ A1 A2 : rel, R ++> option_rel R). 547 | Proof. 548 | exact @Some_rel_def. 549 | Qed. 550 | 551 | Global Instance None_rel: 552 | Monotonic (@None) (forallr R, option_rel R). 553 | Proof. 554 | exact @None_rel_def. 555 | Qed. 556 | 557 | Global Instance option_subrel {A1 A2}: 558 | Monotonic (@option_rel A1 A2) (subrel ++> subrel). 559 | Proof. 560 | intros RA1 RA2 HRA. 561 | intros x1 x2 Hx. 562 | destruct Hx; constructor; eauto. 563 | Qed. 564 | 565 | Global Instance option_subrel_params: 566 | Params (@option_rel) 3 := { }. 567 | 568 | Lemma option_rel_refl `(HR: Reflexive): 569 | Reflexive (option_rel R). 570 | Proof. 571 | intros [x | ]; constructor; reflexivity. 572 | Qed. 573 | 574 | Global Hint Extern 1 (Reflexive (option_rel _)) => 575 | eapply option_rel_refl : typeclass_instances. 576 | 577 | Global Instance option_map_rel: 578 | Monotonic 579 | (@option_map) 580 | (forallr RA, forallr RB, (RA ++> RB) ++> option_rel RA ++> option_rel RB). 581 | Proof. 582 | intros A1 A2 RA B1 B2 RB f g Hfg x y Hxy. 583 | destruct Hxy; constructor; eauto. 584 | Qed. 585 | 586 | (** XXX: This does not fit any of our existing patterns, we should 587 | drop it for consistency or introduce a new convention and generalize 588 | this kind of lemmas. *) 589 | 590 | Lemma option_rel_some_inv A B (R: rel A B) (x: option A) (y: option B) (a: A): 591 | option_rel R x y -> 592 | x = Some a -> 593 | exists b, 594 | y = Some b /\ 595 | R a b. 596 | Proof. 597 | destruct 1; inversion 1; subst; eauto. 598 | Qed. 599 | 600 | (** *** Lists *) 601 | 602 | Inductive list_rel {A1 A2} (R: rel A1 A2): rel (list A1) (list A2) := 603 | | nil_rel_def: (list_rel R) (@nil A1) (@nil A2) 604 | | cons_rel_def: (R ++> list_rel R ++> list_rel R)%rel (@cons A1) (@cons A2). 605 | 606 | Global Instance nil_rel: 607 | Monotonic (@nil) (forallr R, list_rel R). 608 | Proof. 609 | exact @nil_rel_def. 610 | Qed. 611 | 612 | Global Instance cons_rel: 613 | Monotonic (@cons) (forallr R, R ++> list_rel R ++> list_rel R). 614 | Proof. 615 | exact @cons_rel_def. 616 | Qed. 617 | 618 | Global Instance list_subrel {A1 A2}: 619 | Monotonic (@list_rel A1 A2) (subrel ++> subrel). 620 | Proof. 621 | intros R S HRS x y. 622 | red in HRS. 623 | induction 1; constructor; eauto. 624 | Qed. 625 | 626 | Global Instance list_subrel_params: 627 | Params (@list_rel) 3 := { }. 628 | 629 | Lemma list_rel_refl `(HR: Reflexive): 630 | Reflexive (list_rel R). 631 | Proof. 632 | intros l. 633 | induction l; constructor; eauto. 634 | Qed. 635 | 636 | Global Hint Extern 1 (Reflexive (list_rel _)) => 637 | eapply list_rel_refl : typeclass_instances. 638 | 639 | Lemma list_rel_corefl `(HR: Coreflexive): 640 | Coreflexive (list_rel R). 641 | Proof. 642 | intros l1 l2 Hl. 643 | induction Hl as [ | x1 x2 Hx l1 l2 Hl IHl]; 644 | f_equal; eauto using coreflexivity. 645 | Qed. 646 | 647 | Global Hint Extern 1 (Coreflexive (list_rel _)) => 648 | eapply list_rel_corefl : typeclass_instances. 649 | 650 | Lemma list_rel_sym `(HR: Symmetric): 651 | Symmetric (list_rel R). 652 | Proof. 653 | intros l1 l2 Hl. 654 | induction Hl; constructor; eauto. 655 | Qed. 656 | 657 | Global Hint Extern 1 (Symmetric (list_rel _)) => 658 | eapply list_rel_sym : typeclass_instances. 659 | 660 | Lemma list_rel_trans `(HR: Transitive): 661 | Transitive (list_rel R). 662 | Proof. 663 | intros l1 l2 l3 Hl12 Hl23. 664 | revert l3 Hl23. 665 | induction Hl12; inversion 1; constructor; eauto. 666 | Qed. 667 | 668 | Global Hint Extern 1 (Transitive (list_rel _)) => 669 | eapply list_rel_trans : typeclass_instances. 670 | 671 | Global Instance length_rel: 672 | Monotonic 673 | (@length) 674 | (forallr R : rel, list_rel R ++> eq). 675 | Proof. 676 | induction 1; simpl; congruence. 677 | Qed. 678 | 679 | Global Instance app_rel: 680 | Monotonic 681 | (@app) 682 | (forallr R : rel, list_rel R ++> list_rel R ++> list_rel R). 683 | Proof. 684 | intros A1 A2 R l1 l2 Hl. 685 | induction Hl as [ | x1 x2 Hx l1 l2 Hl IHl]; simpl. 686 | - firstorder. 687 | - constructor; eauto. 688 | Qed. 689 | 690 | Global Instance map_rel: 691 | Monotonic 692 | (@map) 693 | (forallr RA, forallr RB, (RA ++> RB) ++> list_rel RA ++> list_rel RB). 694 | Proof. 695 | intros A1 A2 RA B1 B2 RB f g Hfg l1 l2 Hl. 696 | induction Hl; constructor; eauto. 697 | Qed. 698 | 699 | Global Instance fold_right_rel: 700 | Monotonic 701 | (@fold_right) 702 | (forallr RA, forallr RB, (RB ++> RA ++> RA) ++> RA ++> list_rel RB ++> RA). 703 | Proof. 704 | intros A1 A2 RA B1 B2 RB f g Hfg a1 a2 Ha l1 l2 Hl. 705 | induction Hl; simpl; eauto. 706 | eapply Hfg; eauto. 707 | Qed. 708 | 709 | Global Instance fold_left_rel: 710 | Monotonic 711 | (@fold_left) 712 | (forallr RA, forallr RB, (RA ++> RB ++> RA) ++> list_rel RB ++> RA ++> RA). 713 | Proof. 714 | intros A1 A2 RA B1 B2 RB f g Hfg l1 l2 Hl. 715 | induction Hl; simpl. 716 | - firstorder. 717 | - intros a1 a2 Ha. 718 | eapply IHHl. 719 | eapply Hfg; assumption. 720 | Qed. 721 | 722 | (** ** Monotonicity of logical connectives *) 723 | 724 | Lemma fold_impl_rstep (A B: Prop): 725 | RStep (impl A B) (A -> B). 726 | Proof. 727 | firstorder. 728 | Qed. 729 | 730 | Global Hint Extern 1 (RStep _ (_ -> _)) => 731 | eapply fold_impl_rstep : typeclass_instances. 732 | 733 | Global Instance all_monotonic {A}: 734 | Monotonic (@all A) ((- ==> impl) ++> impl). 735 | Proof. 736 | intros P Q HPQ H x. 737 | apply HPQ. 738 | apply H. 739 | Qed. 740 | 741 | Global Instance all_monotonic_params: 742 | Params (@all) 1 := { }. 743 | 744 | Global Instance ex_monotonic A: 745 | Monotonic (@ex A) ((- ==> impl) ++> impl). 746 | Proof. 747 | intros P Q HPQ [x Hx]. 748 | exists x. 749 | apply HPQ. 750 | assumption. 751 | Qed. 752 | 753 | Global Instance ex_monotonic_params: 754 | Params (@ex) 1 := { }. 755 | 756 | Global Instance and_monotonic: 757 | Monotonic (@and) (impl ++> impl ++> impl). 758 | Proof. 759 | intros P1 P2 HP Q1 Q2 HQ [HP1 HQ1]. 760 | eauto. 761 | Qed. 762 | 763 | Global Instance or_monotonic: 764 | Monotonic (@or) (impl ++> impl ++> impl). 765 | Proof. 766 | intros P1 P2 HP Q1 Q2 HQ [HP1|HQ1]; 767 | eauto. 768 | Qed. 769 | -------------------------------------------------------------------------------- /Transport.v: -------------------------------------------------------------------------------- 1 | Require Export Monotonicity. 2 | Require Import KLR. 3 | 4 | (** ** The [transport] tactic *) 5 | 6 | (** Often, we know that a number of terms are related, and need to 7 | transport hypotheses built out the the left-hand sides into ones 8 | with similar shapes, but built out of the right-hand sides. For 9 | instance, in cirumstances where [rauto] can establish 10 | [option_rel R x y], we will want to turn a hypothesis of the form 11 | [x = Some a] into one of the form [y = Some b], and remember that 12 | [R a b]. This is the role of the [transport] tactic. *) 13 | 14 | (** *** Transportable hypotheses *) 15 | 16 | (* To make the tactic extensible, we use the following typeclass to 17 | declare patterns of relations and hypothesis shapes that can be 18 | handled in this way. *) 19 | 20 | Class Transport {A B} (R: rel A B) (a: A) (b: B) (PA: Prop) (PB: Prop) := 21 | transport: R a b -> PA -> PB. 22 | 23 | (** One stereotypical example is [option_le]: the [Transport] instance 24 | defined in the [OptionRel] library allows us to transport hypothese 25 | of the form [x = Some a] into hypotheses of the form [y = Some b], 26 | with [a] and [b] related by [R] whenever [x] and [y] are related by 27 | [option_le R], and similarly for [None] and [option_rel]. 28 | 29 | Other instances are declared below. *) 30 | 31 | Global Instance prod_rel_transport_eq_pair {A1 B1 A2 B2} (R1: rel A1 B1) (R2: rel A2 B2) x y a1 a2: 32 | Transport (prod_rel R1 R2) x y (x = (a1, a2)) (exists b1 b2, y = (b1, b2) /\ R1 a1 b1 /\ R2 a2 b2). 33 | Proof. 34 | intros [Hxy1 Hxy2] Hx. 35 | destruct y. 36 | subst. 37 | simpl in *. 38 | eauto. 39 | Qed. 40 | 41 | (** For [set_le] the situation is slightly more involved, for two 42 | reasons. First, a regular [eapply set_le_transport] fails to unify 43 | the parameter [B] of [Transport] against the [_ -> Prop] provied by 44 | the instance below. This can be worked around by pre-unifying that 45 | specific parameter. Second, because [set_le_transport] is 46 | potentially applicable to virtually any hypothesis (since there is 47 | no strongly distinguishing syntactic format in our target 48 | hypotheses), it makes [transport_hyps] very slow. 49 | 50 | To address this, we explicitely register hints based on the 51 | [set_le_transport] tactic, which looks for "keywords" before doing 52 | anything, then applies the lemma with the required unification 53 | preparation. For example, [set_le_transport] is used in the 54 | following way in the [SimClight] library: 55 | << 56 | Hint Extern 1 (Transport _ _ _ _ _) => 57 | set_le_transport @assign_loc : typeclass_instances. 58 | >> 59 | Note that it's necessary to use [@] because [assign_loc] is 60 | parametrized by typeclasses, and we want to avoid undue 61 | specialization at hint registration time. *) 62 | 63 | Lemma set_le_transport {A B} (R: rel A B) sA sB a: 64 | Transport (set_le R) sA sB (sA a) (exists b, sB b /\ R a b). 65 | Proof. 66 | intros HsAB Ha. 67 | edestruct HsAB; eauto. 68 | Qed. 69 | 70 | Ltac set_le_transport keyword := 71 | lazymatch goal with 72 | | |- @Transport ?A ?B ?R ?a ?b ?PA ?PB => 73 | lazymatch PA with 74 | | context [keyword] => 75 | let Xv := fresh "X" in evar (Xv: Type); 76 | let X := eval red in Xv in clear Xv; 77 | unify B (X -> Prop); 78 | eapply set_le_transport 79 | end 80 | end. 81 | 82 | (** We defined a few more relation patterns based on [set_le] and 83 | [rel_curry], with a similar strategy. *) 84 | 85 | Lemma rel_curry_set_le_transport {A1 A2 B1 B2} R sA sB (a1: A1) (a2: A2): 86 | Transport (% set_le R) sA sB 87 | (sA a1 a2) 88 | (exists (b1: B1) (b2: B2), sB b1 b2 /\ R (a1, a2) (b1, b2)). 89 | Proof. 90 | intros HsAB Ha. 91 | destruct (HsAB (a1, a2)) as ([b1 b2] & Hb & Hab); eauto. 92 | Qed. 93 | 94 | Ltac rel_curry_set_le_transport keyword := 95 | lazymatch goal with 96 | | |- @Transport ?A ?B ?R ?a ?b ?PA ?PB => 97 | lazymatch PA with 98 | | context [keyword] => 99 | let Xv := fresh "X" in evar (Xv: Type); 100 | let X := eval red in Xv in clear Xv; 101 | let Yv := fresh "Y" in evar (Yv: Type); 102 | let Y := eval red in Yv in clear Yv; 103 | unify B (X -> Y -> Prop); 104 | eapply rel_curry_set_le_transport 105 | end 106 | end. 107 | 108 | Lemma rel_curry2_set_le_transport {A1 A2 A3 B1 B2 B3} R sA sB (a1: A1) (a2: A2) (a3: A3): 109 | Transport (% % set_le R) sA sB 110 | (sA a1 a2 a3) 111 | (exists (b1: B1) (b2: B2) (b3: B3), sB b1 b2 b3 /\ R (a1, a2, a3) (b1, b2, b3)). 112 | Proof. 113 | intros HsAB Ha. 114 | destruct (HsAB (a1, a2, a3)) as ([[b1 b2] b3] & Hb & Hab); eauto. 115 | Qed. 116 | 117 | Ltac rel_curry2_set_le_transport keyword := 118 | lazymatch goal with 119 | | |- @Transport ?A ?B ?R ?a ?b ?PA ?PB => 120 | lazymatch PA with 121 | | context [keyword] => 122 | let Xv := fresh "X" in evar (Xv: Type); 123 | let X := eval red in Xv in clear Xv; 124 | let Yv := fresh "Y" in evar (Yv: Type); 125 | let Y := eval red in Yv in clear Yv; 126 | let Zv := fresh "Y" in evar (Yv: Type); 127 | let Z := eval red in Yv in clear Yv; 128 | unify B (X -> Y -> Z -> Prop); 129 | eapply rel_curry2_set_le_transport 130 | end 131 | end. 132 | 133 | (** The situation for [impl] is similar: since this transport instance 134 | can apply to pretty much anything, we need to register it on a 135 | case-by-case basis. Here is an example used in the CertiKOS proof 136 | for hypotheses of the form [writable_block ge b]: 137 | << 138 | Hint Extern 10 (Transport _ _ _ (writable_block _ _) _) => 139 | eapply impl_transport : typeclass_instances. 140 | >> *) 141 | 142 | Lemma impl_transport P Q: 143 | Transport impl P Q P Q. 144 | Proof. 145 | firstorder. 146 | Qed. 147 | 148 | (*** *** The actual tactic *) 149 | 150 | (** Often, the target hypotheses declared using the [Transport] class 151 | have existential quantifiers, and need to be broken up to get to the 152 | actual relational hypotheses we're interested in. The [split_hyp] 153 | tactic does that. We also use that opportunity to split up 154 | [prod_rel] and [rel_incr] when appropriate. For [prod_rel] this is 155 | especially useful when [rel_curry] is involved. Likewise we 156 | generally want to split up [rel_incr] as soon as possible, so that 157 | the existentially quantified world therein can be used to 158 | instantiate evars that have been spawned from that point on. 159 | 160 | As a generally useful complement, the [split_hyps] tactic applies 161 | the same process to all hypotheses. *) 162 | 163 | Ltac split_hyp H := 164 | lazymatch type of H with 165 | | ex _ => 166 | destruct H as [? H]; 167 | split_hyp H 168 | | _ /\ _ => 169 | let H1 := fresh in 170 | let H2 := fresh in 171 | destruct H as [H1 H2]; 172 | split_hyp H1; 173 | split_hyp H2 174 | | prod_rel ?Rx ?Ry (?x1, ?y1) (?x2, ?y2) => 175 | change (Rx x1 x2 /\ Ry y1 y2) in H; 176 | split_hyp H 177 | | rel_incr ?acc ?R ?w ?x ?y => 178 | let w' := fresh w "'" in 179 | let Hw' := fresh "H" w' in 180 | destruct H as (w' & Hw' & H); 181 | split_hyp H 182 | | klr_diam ?l ?R ?w ?x ?y => 183 | let w' := fresh w "'" in 184 | let Hw' := fresh "H" w' in 185 | destruct H as (w' & Hw' & H); 186 | split_hyp H 187 | | prod_klr ?Rx ?Ry ?w (?x1, ?y1) (?x2, ?y2) => 188 | change (Rx w x1 x2 /\ Ry w y1 y2) in H; 189 | split_hyp H 190 | | _ => 191 | idtac 192 | end. 193 | 194 | Ltac split_hyps := 195 | repeat 196 | lazymatch goal with 197 | | H: ex _ |- _ => 198 | destruct H 199 | | H: _ /\ _ |- _ => 200 | destruct H 201 | | H: prod_rel ?Rx ?Ry (?x1, ?y1) (?x2, ?y2) |- _ => 202 | change (Rx x1 x2 /\ Ry y1 y2) in H 203 | | H: klr_diam ?R ?w ?x ?y |- _ => 204 | let w' := fresh w "'" in 205 | let Hw' := fresh "H" w' in 206 | destruct H as (w' & Hw' & H) 207 | | H: prod_klr ?Rx ?Ry ?w (?x1, ?y1) (?x2, ?y2) |- _ => 208 | change (Rx w x1 x2 /\ Ry w y1 y2) in H 209 | end. 210 | 211 | (** We're now ready to defined the [transport] tactic, which 212 | essentially looks up a [Transport] instance and applies it the the 213 | hypothesis to be transported, using [RAuto] to solve the relational 214 | subgoal. Note that the relation and right-hand side will usually 215 | contain existential variables, but the proof search can hopefully 216 | proceed by following the structure of the left-hand side. Once the 217 | [Transport] instance has been applied, we use [split_hyp] on the 218 | modified hypothesis as post-processing. 219 | 220 | Because we have many open-ended evars involved, it is easy for 221 | trivial relations such as [eq] to kick in and sabotage us, 222 | overriding the more interesting properties that we actually want to 223 | use and preventing any progress from actually being made. In order 224 | to attenuate this issue, we require that [RAuto] yield unconvertible 225 | related elements. 226 | 227 | Another pitfall we want to avoid is illustrated by the [option_rel] 228 | case. When we have a hypothesis of the form [H: m = Some a], but no 229 | interesting way to relate [m] to something else, then the search for 230 | [(option_rel ?R) m ?n] can end up using [H] itself (because 231 | [option_rel eq] is reflexive, hence [subrel_eq] applies). To prevent 232 | such cases we need to make sure that the hypothesis being 233 | transported is not used to discharge the relational premise, and so 234 | we revert it into the conclusing before we proceed, and use the 235 | following lemma to work on a goal of that form directly. *) 236 | 237 | Lemma transport_in_goal `{Transport} `{!RAuto (R a b)} `{!Unconvertible a b}: 238 | forall (Q: Prop), (PB -> Q) -> (PA -> Q). 239 | Proof. 240 | firstorder. 241 | Qed. 242 | 243 | Ltac transport H := 244 | revert H; 245 | lazymatch goal with 246 | | |- ?PA -> ?Q => 247 | apply (transport_in_goal (PA:=PA) Q) 248 | end; 249 | intro H; 250 | split_hyp H. 251 | 252 | (** Again we provide a tactic which attempts to transport all 253 | hypotheses. Notice that earlier transportations may provide new 254 | hypotheses making later transportations possible. Hence it would be 255 | hard to provide a much more effective tactic, even though this one 256 | may retry failing transportations many times. *) 257 | 258 | Ltac transport_hyps := 259 | repeat 260 | match goal with 261 | | H: _ |- _ => transport H 262 | end. 263 | -------------------------------------------------------------------------------- /_CoqProject: -------------------------------------------------------------------------------- 1 | -R . coqrel 2 | 3 | Delay.v 4 | RelDefinitions.v 5 | RelClasses.v 6 | RelOperators.v 7 | Relators.v 8 | Monotonicity.v 9 | RDestruct.v 10 | MorphismsInterop.v 11 | Transport.v 12 | PreOrderTactic.v 13 | LogicalRelations.v 14 | 15 | BoolRel.v 16 | OptionRel.v 17 | KLR.v 18 | 19 | LogicalRelationsTests.v 20 | -------------------------------------------------------------------------------- /configure: -------------------------------------------------------------------------------- 1 | #!/bin/sh 2 | coq_makefile -f _CoqProject -o Makefile 3 | -------------------------------------------------------------------------------- /index.html: -------------------------------------------------------------------------------- 1 | 2 | 3 | 4 | 5 | Coqrel: a binary logical relations library for the Coq proof assistant 6 | 7 | 13 | 14 | 15 |

16 | Coqrel: binary logical relations for Coq 17 |

18 | 19 |

Introduction

20 |

21 | Coqrel is a library for the Coq proof assistant 22 | which mechanizes techniques associated with binary logical relations. 23 | It provides a unified syntax and methodology for 24 | building relations between complex objects, 25 | and an extensible tactic library for solving relational goals. 26 |

27 | 28 |

Download

29 |

30 | The code is available 31 | on github. 32 | Beware: Coqrel is a work in progress, 33 | and its exact interface has not been finalized yet. 34 | With your help and feedback, 35 | I will do my best to make it as usable as possible, 36 | but I make no promise to preserve backward compatibility 37 | until I release the first stable version. 38 |

39 | 40 |

Documentation

41 |

42 | I presented Coqrel at the CoqPL'16 workshop; an 43 | extended abstract 44 | is available which provides a good high-level overview. 45 | The following paper may be philosophically useful as well: 46 | 47 | Logical Relations and Parametricity - 48 | A Reynolds Programme for Category Theory and Programming languages 49 | by C. Hermida, U. S. Reddy and E. P. Robinson. WACT 2013. 50 |

51 |

52 | While documentation is scarce, 53 | the code is well-commented and should be readable. 54 | You can browse the Coqdoc 55 | table of contents and 56 | index here. 57 | The top-level file 58 | LogicalRelations 59 | is a good place to start. 60 | Coqrel uses typeclasses extensively, 61 | both to make the system extensible by the user, 62 | and to enable full backtracking between the different components. 63 | As a result, 64 | some familiarity with the typeclass system of Coq 65 | will be valuable to understand the code. 66 |

67 | 68 |

Contact

69 |

70 | I encourage you to use github's 71 | issue tracker 72 | to report any problems you encounter with Coqrel 73 | or suggest improvements. 74 | For general questions you can also email me directly at 75 | jeremie.koenig@yale.edu; 76 | I'm always happy to hear from potential users of Coqrel! 77 |

78 | 79 |

80 | -- Jérémie Koenig 81 |

82 | 83 | 84 | --------------------------------------------------------------------------------