aboutsummaryrefslogtreecommitdiffhomepage
path: root/theories/Setoids
diff options
context:
space:
mode:
authorGravatar sacerdot <sacerdot@85f007b7-540e-0410-9357-904b9bb8a0f7>2004-09-29 13:06:57 +0000
committerGravatar sacerdot <sacerdot@85f007b7-540e-0410-9357-904b9bb8a0f7>2004-09-29 13:06:57 +0000
commita3ebf0bc16223643a7cdf5da69b4c6a319bad73c (patch)
tree7cc42c68030f0c8177773b8956ccd9266f06d38f /theories/Setoids
parentbef5058398fc59c4b010977967d9a58730e62e1b (diff)
impl is a reflexive relation (it used to be areflexive).
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@6154 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'theories/Setoids')
-rw-r--r--theories/Setoids/Setoid.v6
1 files changed, 5 insertions, 1 deletions
diff --git a/theories/Setoids/Setoid.v b/theories/Setoids/Setoid.v
index 7d8196a72..1bf729055 100644
--- a/theories/Setoids/Setoid.v
+++ b/theories/Setoids/Setoid.v
@@ -608,9 +608,13 @@ Add Morphism not : Not_Morphism.
tauto.
Qed.
+Theorem impl_refl: reflexive _ impl.
+ hnf; unfold impl; tauto.
+Qed.
+
(* THE ASYMMETRIC AREFLEXIVE RELATION impl WITH A FEW MORPHISMS *)
-Add Relation Prop impl.
+Add Relation Prop impl reflexivity proved by impl_refl.
Add Morphism impl with signature impl --> impl ++> impl as Impl_Morphism2.
unfold impl; tauto.