aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--tactics/setoid_replace.ml2
-rw-r--r--theories/Setoids/Setoid.v6
2 files changed, 6 insertions, 2 deletions
diff --git a/tactics/setoid_replace.ml b/tactics/setoid_replace.ml
index 14a81311b..77a5f2116 100644
--- a/tactics/setoid_replace.ml
+++ b/tactics/setoid_replace.ml
@@ -194,7 +194,7 @@ let coq_prop_relation2 =
Relation {
rel_a = mkProp ;
rel_aeq = Lazy.force coq_impl ;
- rel_refl = None;
+ rel_refl = Some (constant ["Setoid"] "impl_refl") ;
rel_sym = None
})
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.