summaryrefslogtreecommitdiff
path: root/theories/Classes/RelationClasses.v
diff options
context:
space:
mode:
Diffstat (limited to 'theories/Classes/RelationClasses.v')
-rw-r--r--theories/Classes/RelationClasses.v432
1 files changed, 228 insertions, 204 deletions
diff --git a/theories/Classes/RelationClasses.v b/theories/Classes/RelationClasses.v
index 5c4dd532..1a40e5d5 100644
--- a/theories/Classes/RelationClasses.v
+++ b/theories/Classes/RelationClasses.v
@@ -1,7 +1,7 @@
(* -*- coding: utf-8 -*- *)
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -20,43 +20,191 @@ Require Import Coq.Program.Basics.
Require Import Coq.Program.Tactics.
Require Import Coq.Relations.Relation_Definitions.
-(** We allow to unfold the [relation] definition while doing morphism search. *)
-
-Notation inverse R := (flip (R:relation _) : relation _).
-
-Definition complement {A} (R : relation A) : relation A := fun x y => R x y -> False.
-
-(** Opaque for proof-search. *)
-Typeclasses Opaque complement.
-
-(** These are convertible. *)
-
-Lemma complement_inverse : forall A (R : relation A), complement (inverse R) = inverse (complement R).
-Proof. reflexivity. Qed.
+Generalizable Variables A B C D R S T U l eqA eqB eqC eqD.
-(** We rebind relations in separate classes to be able to overload each proof. *)
+(** We allow to unfold the [relation] definition while doing morphism search. *)
-Set Implicit Arguments.
-Unset Strict Implicit.
+Section Defs.
+ Context {A : Type}.
+
+ (** We rebind relational properties in separate classes to be able to overload each proof. *)
+
+ Class Reflexive (R : relation A) :=
+ reflexivity : forall x : A, R x x.
+
+ Definition complement (R : relation A) : relation A := fun x y => R x y -> False.
+
+ (** Opaque for proof-search. *)
+ Typeclasses Opaque complement.
+
+ (** These are convertible. *)
+ Lemma complement_inverse R : complement (flip R) = flip (complement R).
+ Proof. reflexivity. Qed.
+
+ Class Irreflexive (R : relation A) :=
+ irreflexivity : Reflexive (complement R).
+
+ Class Symmetric (R : relation A) :=
+ symmetry : forall {x y}, R x y -> R y x.
+
+ Class Asymmetric (R : relation A) :=
+ asymmetry : forall {x y}, R x y -> R y x -> False.
+
+ Class Transitive (R : relation A) :=
+ transitivity : forall {x y z}, R x y -> R y z -> R x z.
+
+ (** Various combinations of reflexivity, symmetry and transitivity. *)
+
+ (** A [PreOrder] is both Reflexive and Transitive. *)
+
+ Class PreOrder (R : relation A) : Prop := {
+ PreOrder_Reflexive :> Reflexive R | 2 ;
+ PreOrder_Transitive :> Transitive R | 2 }.
+
+ (** A [StrictOrder] is both Irreflexive and Transitive. *)
+
+ Class StrictOrder (R : relation A) : Prop := {
+ StrictOrder_Irreflexive :> Irreflexive R ;
+ StrictOrder_Transitive :> Transitive R }.
+
+ (** By definition, a strict order is also asymmetric *)
+ Global Instance StrictOrder_Asymmetric `(StrictOrder R) : Asymmetric R.
+ Proof. firstorder. Qed.
+
+ (** A partial equivalence relation is Symmetric and Transitive. *)
+
+ Class PER (R : relation A) : Prop := {
+ PER_Symmetric :> Symmetric R | 3 ;
+ PER_Transitive :> Transitive R | 3 }.
+
+ (** Equivalence relations. *)
+
+ Class Equivalence (R : relation A) : Prop := {
+ Equivalence_Reflexive :> Reflexive R ;
+ Equivalence_Symmetric :> Symmetric R ;
+ Equivalence_Transitive :> Transitive R }.
+
+ (** An Equivalence is a PER plus reflexivity. *)
+
+ Global Instance Equivalence_PER {R} `(E:Equivalence R) : PER R | 10 :=
+ { }.
+
+ (** We can now define antisymmetry w.r.t. an equivalence relation on the carrier. *)
+
+ Class Antisymmetric eqA `{equ : Equivalence eqA} (R : relation A) :=
+ antisymmetry : forall {x y}, R x y -> R y x -> eqA x y.
+
+ Class subrelation (R R' : relation A) : Prop :=
+ is_subrelation : forall {x y}, R x y -> R' x y.
+
+ (** Any symmetric relation is equal to its inverse. *)
+
+ Lemma subrelation_symmetric R `(Symmetric R) : subrelation (flip R) R.
+ Proof. hnf. intros. red in H0. apply symmetry. assumption. Qed.
+
+ Section flip.
+
+ Lemma flip_Reflexive `{Reflexive R} : Reflexive (flip R).
+ Proof. tauto. Qed.
+
+ Program Definition flip_Irreflexive `(Irreflexive R) : Irreflexive (flip R) :=
+ irreflexivity (R:=R).
+
+ Program Definition flip_Symmetric `(Symmetric R) : Symmetric (flip R) :=
+ fun x y H => symmetry (R:=R) H.
+
+ Program Definition flip_Asymmetric `(Asymmetric R) : Asymmetric (flip R) :=
+ fun x y H H' => asymmetry (R:=R) H H'.
+
+ Program Definition flip_Transitive `(Transitive R) : Transitive (flip R) :=
+ fun x y z H H' => transitivity (R:=R) H' H.
+
+ Program Definition flip_Antisymmetric `(Antisymmetric eqA R) :
+ Antisymmetric eqA (flip R).
+ Proof. firstorder. Qed.
+
+ (** Inversing the larger structures *)
+
+ Lemma flip_PreOrder `(PreOrder R) : PreOrder (flip R).
+ Proof. firstorder. Qed.
+
+ Lemma flip_StrictOrder `(StrictOrder R) : StrictOrder (flip R).
+ Proof. firstorder. Qed.
+
+ Lemma flip_PER `(PER R) : PER (flip R).
+ Proof. firstorder. Qed.
+
+ Lemma flip_Equivalence `(Equivalence R) : Equivalence (flip R).
+ Proof. firstorder. Qed.
+
+ End flip.
+
+ Section complement.
+
+ Definition complement_Irreflexive `(Reflexive R)
+ : Irreflexive (complement R).
+ Proof. firstorder. Qed.
+
+ Definition complement_Symmetric `(Symmetric R) : Symmetric (complement R).
+ Proof. firstorder. Qed.
+ End complement.
+
+
+ (** Rewrite relation on a given support: declares a relation as a rewrite
+ relation for use by the generalized rewriting tactic.
+ It helps choosing if a rewrite should be handled
+ by the generalized or the regular rewriting tactic using leibniz equality.
+ Users can declare an [RewriteRelation A RA] anywhere to declare default
+ relations. This is also done automatically by the [Declare Relation A RA]
+ commands. *)
-Class Reflexive {A} (R : relation A) :=
- reflexivity : forall x, R x x.
+ Class RewriteRelation (RA : relation A).
-Class Irreflexive {A} (R : relation A) :=
- irreflexivity : Reflexive (complement R).
+ (** Any [Equivalence] declared in the context is automatically considered
+ a rewrite relation. *)
+
+ Global Instance equivalence_rewrite_relation `(Equivalence eqA) : RewriteRelation eqA.
+
+ (** Leibniz equality. *)
+ Section Leibniz.
+ Global Instance eq_Reflexive : Reflexive (@eq A) := @eq_refl A.
+ Global Instance eq_Symmetric : Symmetric (@eq A) := @eq_sym A.
+ Global Instance eq_Transitive : Transitive (@eq A) := @eq_trans A.
+
+ (** Leibinz equality [eq] is an equivalence relation.
+ The instance has low priority as it is always applicable
+ if only the type is constrained. *)
+
+ Global Program Instance eq_equivalence : Equivalence (@eq A) | 10.
+ End Leibniz.
+
+End Defs.
+
+(** Default rewrite relations handled by [setoid_rewrite]. *)
+Instance: RewriteRelation impl.
+Instance: RewriteRelation iff.
+(** Hints to drive the typeclass resolution avoiding loops
+ due to the use of full unification. *)
Hint Extern 1 (Reflexive (complement _)) => class_apply @irreflexivity : typeclass_instances.
+Hint Extern 3 (Symmetric (complement _)) => class_apply complement_Symmetric : typeclass_instances.
+Hint Extern 3 (Irreflexive (complement _)) => class_apply complement_Irreflexive : typeclass_instances.
-Class Symmetric {A} (R : relation A) :=
- symmetry : forall x y, R x y -> R y x.
+Hint Extern 3 (Reflexive (flip _)) => apply flip_Reflexive : typeclass_instances.
+Hint Extern 3 (Irreflexive (flip _)) => class_apply flip_Irreflexive : typeclass_instances.
+Hint Extern 3 (Symmetric (flip _)) => class_apply flip_Symmetric : typeclass_instances.
+Hint Extern 3 (Asymmetric (flip _)) => class_apply flip_Asymmetric : typeclass_instances.
+Hint Extern 3 (Antisymmetric (flip _)) => class_apply flip_Antisymmetric : typeclass_instances.
+Hint Extern 3 (Transitive (flip _)) => class_apply flip_Transitive : typeclass_instances.
+Hint Extern 3 (StrictOrder (flip _)) => class_apply flip_StrictOrder : typeclass_instances.
+Hint Extern 3 (PreOrder (flip _)) => class_apply flip_PreOrder : typeclass_instances.
-Class Asymmetric {A} (R : relation A) :=
- asymmetry : forall x y, R x y -> R y x -> False.
+Hint Extern 4 (subrelation (flip _) _) =>
+ class_apply @subrelation_symmetric : typeclass_instances.
-Class Transitive {A} (R : relation A) :=
- transitivity : forall x y z, R x y -> R y z -> R x z.
+Arguments irreflexivity {A R Irreflexive} [x] _.
-Hint Resolve @irreflexivity : ord.
+Hint Resolve irreflexivity : ord.
Unset Implicit Arguments.
@@ -72,40 +220,6 @@ Hint Extern 4 => solve_relation : relations.
(** We can already dualize all these properties. *)
-Generalizable Variables A B C D R S T U l eqA eqB eqC eqD.
-
-Lemma flip_Reflexive `{Reflexive A R} : Reflexive (flip R).
-Proof. tauto. Qed.
-
-Hint Extern 3 (Reflexive (flip _)) => apply flip_Reflexive : typeclass_instances.
-
-Program Definition flip_Irreflexive `(Irreflexive A R) : Irreflexive (flip R) :=
- irreflexivity (R:=R).
-
-Program Definition flip_Symmetric `(Symmetric A R) : Symmetric (flip R) :=
- fun x y H => symmetry (R:=R) H.
-
-Program Definition flip_Asymmetric `(Asymmetric A R) : Asymmetric (flip R) :=
- fun x y H H' => asymmetry (R:=R) H H'.
-
-Program Definition flip_Transitive `(Transitive A R) : Transitive (flip R) :=
- fun x y z H H' => transitivity (R:=R) H' H.
-
-Hint Extern 3 (Irreflexive (flip _)) => class_apply flip_Irreflexive : typeclass_instances.
-Hint Extern 3 (Symmetric (flip _)) => class_apply flip_Symmetric : typeclass_instances.
-Hint Extern 3 (Asymmetric (flip _)) => class_apply flip_Asymmetric : typeclass_instances.
-Hint Extern 3 (Transitive (flip _)) => class_apply flip_Transitive : typeclass_instances.
-
-Definition Reflexive_complement_Irreflexive `(Reflexive A (R : relation A))
- : Irreflexive (complement R).
-Proof. firstorder. Qed.
-
-Definition complement_Symmetric `(Symmetric A (R : relation A)) : Symmetric (complement R).
-Proof. firstorder. Qed.
-
-Hint Extern 3 (Symmetric (complement _)) => class_apply complement_Symmetric : typeclass_instances.
-Hint Extern 3 (Irreflexive (complement _)) => class_apply Reflexive_complement_Irreflexive : typeclass_instances.
-
(** * Standard instances. *)
Ltac reduce_hyp H :=
@@ -130,7 +244,7 @@ Tactic Notation "apply" "*" constr(t) :=
Ltac simpl_relation :=
unfold flip, impl, arrow ; try reduce ; program_simpl ;
- try ( solve [ intuition ]).
+ try ( solve [ dintuition ]).
Local Obligation Tactic := simpl_relation.
@@ -145,54 +259,6 @@ Instance iff_Reflexive : Reflexive iff := iff_refl.
Instance iff_Symmetric : Symmetric iff := iff_sym.
Instance iff_Transitive : Transitive iff := iff_trans.
-(** Leibniz equality. *)
-
-Instance eq_Reflexive {A} : Reflexive (@eq A) := @eq_refl A.
-Instance eq_Symmetric {A} : Symmetric (@eq A) := @eq_sym A.
-Instance eq_Transitive {A} : Transitive (@eq A) := @eq_trans A.
-
-(** Various combinations of reflexivity, symmetry and transitivity. *)
-
-(** A [PreOrder] is both Reflexive and Transitive. *)
-
-Class PreOrder {A} (R : relation A) : Prop := {
- PreOrder_Reflexive :> Reflexive R | 2 ;
- PreOrder_Transitive :> Transitive R | 2 }.
-
-(** A partial equivalence relation is Symmetric and Transitive. *)
-
-Class PER {A} (R : relation A) : Prop := {
- PER_Symmetric :> Symmetric R | 3 ;
- PER_Transitive :> Transitive R | 3 }.
-
-(** Equivalence relations. *)
-
-Class Equivalence {A} (R : relation A) : Prop := {
- Equivalence_Reflexive :> Reflexive R ;
- Equivalence_Symmetric :> Symmetric R ;
- Equivalence_Transitive :> Transitive R }.
-
-(** An Equivalence is a PER plus reflexivity. *)
-
-Instance Equivalence_PER `(Equivalence A R) : PER R | 10 :=
- { PER_Symmetric := Equivalence_Symmetric ;
- PER_Transitive := Equivalence_Transitive }.
-
-(** We can now define antisymmetry w.r.t. an equivalence relation on the carrier. *)
-
-Class Antisymmetric A eqA `{equ : Equivalence A eqA} (R : relation A) :=
- antisymmetry : forall {x y}, R x y -> R y x -> eqA x y.
-
-Program Definition flip_antiSymmetric `(Antisymmetric A eqA R) :
- Antisymmetric A eqA (flip R).
-Proof. firstorder. Qed.
-
-(** Leibinz equality [eq] is an equivalence relation.
- The instance has low priority as it is always applicable
- if only the type is constrained. *)
-
-Program Instance eq_equivalence : Equivalence (@eq A) | 10.
-
(** Logical equivalence [iff] is an equivalence relation. *)
Program Instance iff_equivalence : Equivalence iff.
@@ -203,9 +269,6 @@ Program Instance iff_equivalence : Equivalence iff.
Local Open Scope list_scope.
-(* Notation " [ ] " := nil : list_scope. *)
-(* Notation " [ x ; .. ; y ] " := (cons x .. (cons y nil) ..) (at level 1) : list_scope. *)
-
(** A compact representation of non-dependent arities, with the codomain singled-out. *)
(* Note, we do not use [list Type] because it imposes unnecessary universe constraints *)
@@ -316,7 +379,8 @@ Notation "∙⊥∙" := false_predicate : predicate_scope.
(** Predicate equivalence is an equivalence, and predicate implication defines a preorder. *)
-Program Instance predicate_equivalence_equivalence : Equivalence (@predicate_equivalence l).
+Program Instance predicate_equivalence_equivalence :
+ Equivalence (@predicate_equivalence l).
Next Obligation.
induction l ; firstorder.
@@ -345,106 +409,66 @@ Program Instance predicate_implication_preorder :
(** We define the various operations which define the algebra on binary relations,
from the general ones. *)
-Definition relation_equivalence {A : Type} : relation (relation A) :=
- @predicate_equivalence (_::_::Tnil).
-
-Class subrelation {A:Type} (R R' : relation A) : Prop :=
- is_subrelation : @predicate_implication (A::A::Tnil) R R'.
-
-Arguments subrelation {A} R R'.
-
-Definition relation_conjunction {A} (R : relation A) (R' : relation A) : relation A :=
- @predicate_intersection (A::A::Tnil) R R'.
-
-Definition relation_disjunction {A} (R : relation A) (R' : relation A) : relation A :=
- @predicate_union (A::A::Tnil) R R'.
-
-(** Relation equivalence is an equivalence, and subrelation defines a partial order. *)
-
-Set Automatic Introduction.
-
-Instance relation_equivalence_equivalence (A : Type) :
- Equivalence (@relation_equivalence A).
-Proof. exact (@predicate_equivalence_equivalence (A::A::Tnil)). Qed.
-
-Instance relation_implication_preorder A : PreOrder (@subrelation A).
-Proof. exact (@predicate_implication_preorder (A::A::Tnil)). Qed.
-
-(** *** Partial Order.
+Section Binary.
+ Context {A : Type}.
+
+ Definition relation_equivalence : relation (relation A) :=
+ @predicate_equivalence (_::_::Tnil).
+
+ Global Instance: RewriteRelation relation_equivalence.
+
+ Definition relation_conjunction (R : relation A) (R' : relation A) : relation A :=
+ @predicate_intersection (A::A::Tnil) R R'.
+
+ Definition relation_disjunction (R : relation A) (R' : relation A) : relation A :=
+ @predicate_union (A::A::Tnil) R R'.
+
+ (** Relation equivalence is an equivalence, and subrelation defines a partial order. *)
+
+ Set Automatic Introduction.
+
+ Global Instance relation_equivalence_equivalence :
+ Equivalence relation_equivalence.
+ Proof. exact (@predicate_equivalence_equivalence (A::A::Tnil)). Qed.
+
+ Global Instance relation_implication_preorder : PreOrder (@subrelation A).
+ Proof. exact (@predicate_implication_preorder (A::A::Tnil)). Qed.
+
+ (** *** Partial Order.
A partial order is a preorder which is additionally antisymmetric.
We give an equivalent definition, up-to an equivalence relation
on the carrier. *)
-Class PartialOrder {A} eqA `{equ : Equivalence A eqA} R `{preo : PreOrder A R} :=
- partial_order_equivalence : relation_equivalence eqA (relation_conjunction R (inverse R)).
+ Class PartialOrder eqA `{equ : Equivalence A eqA} R `{preo : PreOrder A R} :=
+ partial_order_equivalence : relation_equivalence eqA (relation_conjunction R (flip R)).
+
+ (** The equivalence proof is sufficient for proving that [R] must be a
+ morphism for equivalence (see Morphisms). It is also sufficient to
+ show that [R] is antisymmetric w.r.t. [eqA] *)
+
+ Global Instance partial_order_antisym `(PartialOrder eqA R) : ! Antisymmetric A eqA R.
+ Proof with auto.
+ reduce_goal.
+ pose proof partial_order_equivalence as poe. do 3 red in poe.
+ apply <- poe. firstorder.
+ Qed.
-(** The equivalence proof is sufficient for proving that [R] must be a morphism
- for equivalence (see Morphisms).
- It is also sufficient to show that [R] is antisymmetric w.r.t. [eqA] *)
-Instance partial_order_antisym `(PartialOrder A eqA R) : ! Antisymmetric A eqA R.
-Proof with auto.
- reduce_goal.
- pose proof partial_order_equivalence as poe. do 3 red in poe.
- apply <- poe. firstorder.
-Qed.
+ Lemma PartialOrder_inverse `(PartialOrder eqA R) : PartialOrder eqA (flip R).
+ Proof. firstorder. Qed.
+End Binary.
+
+Hint Extern 3 (PartialOrder (flip _)) => class_apply PartialOrder_inverse : typeclass_instances.
(** The partial order defined by subrelation and relation equivalence. *)
Program Instance subrelation_partial_order :
! PartialOrder (relation A) relation_equivalence subrelation.
- Next Obligation.
- Proof.
- unfold relation_equivalence in *. compute; firstorder.
- Qed.
+Next Obligation.
+Proof.
+ unfold relation_equivalence in *. compute; firstorder.
+Qed.
Typeclasses Opaque arrows predicate_implication predicate_equivalence
- relation_equivalence pointwise_lifting.
-
-(** Rewrite relation on a given support: declares a relation as a rewrite
- relation for use by the generalized rewriting tactic.
- It helps choosing if a rewrite should be handled
- by the generalized or the regular rewriting tactic using leibniz equality.
- Users can declare an [RewriteRelation A RA] anywhere to declare default
- relations. This is also done automatically by the [Declare Relation A RA]
- commands. *)
-
-Class RewriteRelation {A : Type} (RA : relation A).
-
-Instance: RewriteRelation impl.
-Instance: RewriteRelation iff.
-Instance: RewriteRelation (@relation_equivalence A).
-
-(** Any [Equivalence] declared in the context is automatically considered
- a rewrite relation. *)
-
-Instance equivalence_rewrite_relation `(Equivalence A eqA) : RewriteRelation eqA.
-
-(** Strict Order *)
-
-Class StrictOrder {A : Type} (R : relation A) : Prop := {
- StrictOrder_Irreflexive :> Irreflexive R ;
- StrictOrder_Transitive :> Transitive R
-}.
-
-Instance StrictOrder_Asymmetric `(StrictOrder A R) : Asymmetric R.
-Proof. firstorder. Qed.
-
-(** Inversing a [StrictOrder] gives another [StrictOrder] *)
-
-Lemma StrictOrder_inverse `(StrictOrder A R) : StrictOrder (inverse R).
-Proof. firstorder. Qed.
-
-(** Same for [PartialOrder]. *)
-
-Lemma PreOrder_inverse `(PreOrder A R) : PreOrder (inverse R).
-Proof. firstorder. Qed.
-
-Hint Extern 3 (StrictOrder (inverse _)) => class_apply StrictOrder_inverse : typeclass_instances.
-Hint Extern 3 (PreOrder (inverse _)) => class_apply PreOrder_inverse : typeclass_instances.
-
-Lemma PartialOrder_inverse `(PartialOrder A eqA R) : PartialOrder eqA (inverse R).
-Proof. firstorder. Qed.
-
-Hint Extern 3 (PartialOrder (inverse _)) => class_apply PartialOrder_inverse : typeclass_instances.
+ relation_equivalence pointwise_lifting.