Built with Alectryon, running Coq+SerAPI v8.10.0+0.7.0. Coq sources are in this panel; goals and messages will appear in the other. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus.
(* -*- coding: utf-8 -*- *)
(************************************************************************)
(* * The Coq Proof Assistant / The Coq Development Team *)
(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
Typeclass-based relations, tactics and standard instances
Require Export Coq.Classes.Init. Require Import Coq.Program.Basics. Require Import Coq.Program.Tactics. Generalizable Variables A B C D R S T U l eqA eqB eqC eqD. Set Universe Polymorphism. Definition crelation (A : Type) := A -> A -> Type. Definition arrow (A B : Type) := A -> B. Definition flip {A B C : Type} (f : A -> B -> C) := fun x y => f y x. Definition iffT (A B : Type) := ((A -> B) * (B -> A))%type.
We allow to unfold the crelation definition while doing morphism search.
Section Defs. Context {A : Type}.
We rebind crelational properties in separate classes to be able to overload each proof.
Class Reflexive (R : crelation A) := reflexivity : forall x : A, R x x. Definition complement (R : crelation A) : crelation A := fun x y => R x y -> False.
Opaque for proof-search.
Typeclasses Opaque complement iffT.
These are convertible.
A:TypeR:A -> A -> Typecomplement (flip R) = flip (complement R)reflexivity. Qed. Class Irreflexive (R : crelation A) := irreflexivity : Reflexive (complement R). Class Symmetric (R : crelation A) := symmetry : forall {x y}, R x y -> R y x. Class Asymmetric (R : crelation A) := asymmetry : forall {x y}, R x y -> (complement R y x : Type). Class Transitive (R : crelation A) := transitivity : forall {x y z}, R x y -> R y z -> R x z.A:TypeR:A -> A -> Typecomplement (flip R) = flip (complement R)
Various combinations of reflexivity, symmetry and transitivity.
A PreOrder is both Reflexive and Transitive.
Class PreOrder (R : crelation A) := { PreOrder_Reflexive :> Reflexive R | 2 ; PreOrder_Transitive :> Transitive R | 2 }.
A StrictOrder is both Irreflexive and Transitive.
Class StrictOrder (R : crelation A) := { StrictOrder_Irreflexive :> Irreflexive R ; StrictOrder_Transitive :> Transitive R }.
By definition, a strict order is also asymmetric
A:TypeR:crelation AH:StrictOrder RAsymmetric Rfirstorder. Qed.A:TypeR:crelation AH:StrictOrder RAsymmetric R
A partial equivalence crelation is Symmetric and Transitive.
Class PER (R : crelation A) := { PER_Symmetric :> Symmetric R | 3 ; PER_Transitive :> Transitive R | 3 }.
Equivalence crelations.
Class Equivalence (R : crelation A) := { Equivalence_Reflexive :> Reflexive R ; Equivalence_Symmetric :> Symmetric R ; Equivalence_Transitive :> Transitive R }.
An Equivalence is a PER plus reflexivity.
Instance Equivalence_PER {R} `(Equivalence R) : PER R | 10 := { PER_Symmetric := Equivalence_Symmetric ; PER_Transitive := Equivalence_Transitive }.
We can now define antisymmetry w.r.t. an equivalence crelation on the carrier.
Class Antisymmetric eqA `{equ : Equivalence eqA} (R : crelation A) := antisymmetry : forall {x y}, R x y -> R y x -> eqA x y. Class subrelation (R R' : crelation A) := is_subrelation : forall {x y}, R x y -> R' x y.
Any symmetric crelation is equal to its inverse.
A:TypeR:crelation AH:Symmetric Rsubrelation (flip R) RA:TypeR:crelation AH:Symmetric Rsubrelation (flip R) RA:TypeR:crelation AH:Symmetric Rforall x y : A, flip R x y -> R x yA:TypeR:crelation AH:Symmetric Rx, y:AH':flip R x yR x yA:TypeR:crelation AH:Symmetric Rx, y:AH':R y xR x yassumption. Qed. Section flip.A:TypeR:crelation AH:Symmetric Rx, y:AH':R y xR y xA:TypeR:crelation AH:Reflexive RReflexive (flip R)tauto. Qed. Definition flip_Irreflexive `(Irreflexive R) : Irreflexive (flip R) := irreflexivity (R:=R). Definition flip_Symmetric `(Symmetric R) : Symmetric (flip R) := fun x y H => symmetry (R:=R) H. Definition flip_Asymmetric `(Asymmetric R) : Asymmetric (flip R) := fun x y H H' => asymmetry (R:=R) H H'. Definition flip_Transitive `(Transitive R) : Transitive (flip R) := fun x y z H H' => transitivity (R:=R) H' H.A:TypeR:crelation AH:Reflexive RReflexive (flip R)A:TypeeqA:crelation Aequ:Equivalence eqAR:crelation AH:Antisymmetric eqA RAntisymmetric eqA (flip R)firstorder. Qed.A:TypeeqA:crelation Aequ:Equivalence eqAR:crelation AH:Antisymmetric eqA RAntisymmetric eqA (flip R)
Inversing the larger structures
A:TypeR:crelation AH:PreOrder RPreOrder (flip R)firstorder. Qed.A:TypeR:crelation AH:PreOrder RPreOrder (flip R)A:TypeR:crelation AH:StrictOrder RStrictOrder (flip R)firstorder. Qed.A:TypeR:crelation AH:StrictOrder RStrictOrder (flip R)A:TypeR:crelation AH:PER RPER (flip R)firstorder. Qed.A:TypeR:crelation AH:PER RPER (flip R)A:TypeR:crelation AH:Equivalence REquivalence (flip R)firstorder. Qed. End flip. Section complement.A:TypeR:crelation AH:Equivalence REquivalence (flip R)A:TypeR:crelation AH:Reflexive RIrreflexive (complement R)firstorder. Qed.A:TypeR:crelation AH:Reflexive RIrreflexive (complement R)A:TypeR:crelation AH:Symmetric RSymmetric (complement R)firstorder. Qed. End complement.A:TypeR:crelation AH:Symmetric RSymmetric (complement R)
Rewrite crelation on a given support: declares a crelation as a rewrite
crelation 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
crelations. This is also done automatically by the Declare Relation A RA
commands.
Class RewriteRelation (RA : crelation A).
Any Equivalence declared in the context is automatically considered
a rewrite crelation.
Instance equivalence_rewrite_crelation `(Equivalence eqA) : RewriteRelation eqA. Defined.
Leibniz equality.
Section Leibniz. Instance eq_Reflexive : Reflexive (@eq A) := @eq_refl A. Instance eq_Symmetric : Symmetric (@eq A) := @eq_sym A. Instance eq_Transitive : Transitive (@eq A) := @eq_trans A.
Leibinz equality eq is an equivalence crelation.
The instance has low priority as it is always applicable
if only the type is constrained.
Instance eq_equivalence : Equivalence (@eq A) | 10. End Leibniz. End Defs.
Default rewrite crelations handled by setoid_rewrite.
Instance: RewriteRelation impl. Defined. Instance: RewriteRelation iff. Defined.
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. 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. Hint Extern 4 (subrelation (flip _) _) => class_apply @subrelation_symmetric : typeclass_instances. Hint Resolve irreflexivity : ord. Unset Implicit Arguments.
A HintDb for crelations.
Ltac solve_crelation := match goal with | [ |- ?R ?x ?x ] => reflexivity | [ H : ?R ?x ?y |- ?R ?y ?x ] => symmetry ; exact H end. Hint Extern 4 => solve_crelation : crelations.
We can already dualize all these properties.
Ltac reduce_hyp H := match type of H with | context [ _ <-> _ ] => fail 1 | _ => red in H ; try reduce_hyp H end. Ltac reduce_goal := match goal with | [ |- _ <-> _ ] => fail 1 | _ => red ; intros ; try reduce_goal end. Tactic Notation "reduce" "in" hyp(Hid) := reduce_hyp Hid. Ltac reduce := reduce_goal. Tactic Notation "apply" "*" constr(t) := first [ refine t | refine (t _) | refine (t _ _) | refine (t _ _ _) | refine (t _ _ _ _) | refine (t _ _ _ _ _) | refine (t _ _ _ _ _ _) | refine (t _ _ _ _ _ _ _) ]. Ltac simpl_crelation := unfold flip, impl, arrow ; try reduce ; program_simpl ; try ( solve [ dintuition ]). Local Obligation Tactic := simpl_crelation.
Logical implication.
Instance impl_Reflexive : Reflexive impl. Instance impl_Transitive : Transitive impl.
Logical equivalence.
Instance iff_Reflexive : Reflexive iff := iff_refl. Instance iff_Symmetric : Symmetric iff := iff_sym. Instance iff_Transitive : Transitive iff := iff_trans.
Logical equivalence iff is an equivalence crelation.
Instance iff_equivalence : Equivalence iff. Instance arrow_Reflexive : Reflexive arrow. Instance arrow_Transitive : Transitive arrow.Reflexive iffTfirstorder. Defined.Reflexive iffTSymmetric iffTfirstorder. Defined.Symmetric iffTTransitive iffTfirstorder. Defined.Transitive iffT
We now develop a generalization of results on crelations for arbitrary predicates.
The resulting theory can be applied to homogeneous binary crelations but also to
arbitrary n-ary predicates.
Local Open Scope list_scope.
A compact representation of non-dependent arities, with the codomain singled-out.
We define the various operations which define the algebra on binary crelations
Section Binary. Context {A : Type}. Definition relation_equivalence : crelation (crelation A) := fun R R' => forall x y, iffT (R x y) (R' x y). Global Instance: RewriteRelation relation_equivalence. Defined. Definition relation_conjunction (R : crelation A) (R' : crelation A) : crelation A := fun x y => prod (R x y) (R' x y). Definition relation_disjunction (R : crelation A) (R' : crelation A) : crelation A := fun x y => sum (R x y) (R' x y).
Relation equivalence is an equivalence, and subrelation defines a partial order.
A:TypeEquivalence relation_equivalenceA:TypeEquivalence relation_equivalenceA:Typeforall (x : crelation A) (x0 y : A), (x x0 y -> x x0 y) * (x x0 y -> x x0 y)A:Typeforall x y : crelation A, (forall x0 y0 : A, (x x0 y0 -> y x0 y0) * (y x0 y0 -> x x0 y0)) -> forall x0 y0 : A, (y x0 y0 -> x x0 y0) * (x x0 y0 -> y x0 y0)A:Typeforall x y z : crelation A, (forall x0 y0 : A, (x x0 y0 -> y x0 y0) * (y x0 y0 -> x x0 y0)) -> (forall x0 y0 : A, (y x0 y0 -> z x0 y0) * (z x0 y0 -> y x0 y0)) -> forall x0 y0 : A, (x x0 y0 -> z x0 y0) * (z x0 y0 -> x x0 y0)firstorder.A:Typeforall (x : crelation A) (x0 y : A), (x x0 y -> x x0 y) * (x x0 y -> x x0 y)firstorder.A:Typeforall x y : crelation A, (forall x0 y0 : A, (x x0 y0 -> y x0 y0) * (y x0 y0 -> x x0 y0)) -> forall x0 y0 : A, (y x0 y0 -> x x0 y0) * (x x0 y0 -> y x0 y0)A:Typeforall x y z : crelation A, (forall x0 y0 : A, (x x0 y0 -> y x0 y0) * (y x0 y0 -> x x0 y0)) -> (forall x0 y0 : A, (y x0 y0 -> z x0 y0) * (z x0 y0 -> y x0 y0)) -> forall x0 y0 : A, (x x0 y0 -> z x0 y0) * (z x0 y0 -> x x0 y0)A:Typex, y, z:crelation AX:forall x1 y1 : A, (x x1 y1 -> y x1 y1) * (y x1 y1 -> x x1 y1)X0:forall x1 y1 : A, (y x1 y1 -> z x1 y1) * (z x1 y1 -> y x1 y1)x0, y0:A((x x0 y0 -> z x0 y0) * (z x0 y0 -> x x0 y0))%typeA:Typex, y, z:crelation Ax0, y0:AX:((x x0 y0 -> y x0 y0) * (y x0 y0 -> x x0 y0))%typeX0:forall x1 y1 : A, (y x1 y1 -> z x1 y1) * (z x1 y1 -> y x1 y1)((x x0 y0 -> z x0 y0) * (z x0 y0 -> x x0 y0))%typefirstorder. Qed.A:Typex, y, z:crelation Ax0, y0:AX:((x x0 y0 -> y x0 y0) * (y x0 y0 -> x x0 y0))%typeX0:((y x0 y0 -> z x0 y0) * (z x0 y0 -> y x0 y0))%type((x x0 y0 -> z x0 y0) * (z x0 y0 -> x x0 y0))%typeA:TypePreOrder subrelationfirstorder. Qed.A:TypePreOrder subrelation
Partial Order.
A partial order is a preorder which is additionally antisymmetric. We give an equivalent definition, up-to an equivalence crelation on the carrier.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
A:TypeeqA:crelation Aequ:Equivalence eqAR:crelation Apreo:PreOrder RH:PartialOrder eqA RAntisymmetric eqA RA:TypeeqA:crelation Aequ:Equivalence eqAR:crelation Apreo:PreOrder RH:PartialOrder eqA RAntisymmetric eqA RA:TypeeqA:crelation Aequ:Equivalence eqAR:crelation Apreo:PreOrder RH:PartialOrder eqA Rx, y:AX:R x yX0:R y xeqA x yfirstorder. Qed.A:TypeeqA:crelation Aequ:Equivalence eqAR:crelation Apreo:PreOrder RH:PartialOrder eqA Rx, y:AX:R x yX0:R y xrelation_conjunction R (flip R) x yA:TypeeqA:crelation Aequ:Equivalence eqAR:crelation Apreo:PreOrder RH:PartialOrder eqA RPartialOrder eqA (flip R)A:TypeeqA:crelation Aequ:Equivalence eqAR:crelation Apreo:PreOrder RH:PartialOrder eqA RPartialOrder eqA (flip R)A:TypeeqA:crelation Aequ:Equivalence eqAR:crelation Apreo:PreOrder RH:PartialOrder eqA Rx, y:AeqA x y -> relation_conjunction (fun x0 y0 : A => R y0 x0) (fun x0 y0 : A => R x0 y0) x yA:TypeeqA:crelation Aequ:Equivalence eqAR:crelation Apreo:PreOrder RH:PartialOrder eqA Rx, y:Arelation_conjunction (fun x0 y0 : A => R y0 x0) (fun x0 y0 : A => R x0 y0) x y -> eqA x yA:TypeeqA:crelation Aequ:Equivalence eqAR:crelation Apreo:PreOrder RH:PartialOrder eqA Rx, y:AeqA x y -> relation_conjunction (fun x0 y0 : A => R y0 x0) (fun x0 y0 : A => R x0 y0) x yA:TypeeqA:crelation Aequ:Equivalence eqAR:crelation Apreo:PreOrder RH:PartialOrder eqA Rx, y:AX:eqA x yrelation_conjunction (fun x0 y0 : A => R y0 x0) (fun x0 y0 : A => R x0 y0) x yA:TypeeqA:crelation Aequ:Equivalence eqAR:crelation Apreo:PreOrder RH:PartialOrder eqA Rx, y:AX:eqA x yeqA y xapply X.A:TypeeqA:crelation Aequ:Equivalence eqAR:crelation Apreo:PreOrder RH:PartialOrder eqA Rx, y:AX:eqA x yeqA x yA:TypeeqA:crelation Aequ:Equivalence eqAR:crelation Apreo:PreOrder RH:PartialOrder eqA Rx, y:Arelation_conjunction (fun x0 y0 : A => R y0 x0) (fun x0 y0 : A => R x0 y0) x y -> eqA x yA:TypeeqA:crelation Aequ:Equivalence eqAR:crelation Apreo:PreOrder RH:PartialOrder eqA Rx, y:AR y x * R x y -> eqA x yA:TypeeqA:crelation Aequ:Equivalence eqAR:crelation Apreo:PreOrder RH:PartialOrder eqA Rx, y:AH1:R y xH2:R x yeqA x yconstructor; assumption. Qed. End Binary. Hint Extern 3 (PartialOrder (flip _)) => class_apply PartialOrder_inverse : typeclass_instances.A:TypeeqA:crelation Aequ:Equivalence eqAR:crelation Apreo:PreOrder RH:PartialOrder eqA Rx, y:AH1:R y xH2:R x yrelation_conjunction R (flip R) x y
The partial order defined by subrelation and crelation equivalence.
(* Program Instance subrelation_partial_order : *) (* ! PartialOrder (crelation A) relation_equivalence subrelation. *) (* Obligation Tactic := idtac. *) (* Next Obligation. *) (* Proof. *) (* intros x. refine (fun x => x). *) (* Qed. *) Typeclasses Opaque relation_equivalence.