aboutsummaryrefslogtreecommitdiffhomepage
path: root/theories/Init/Specif.v
diff options
context:
space:
mode:
authorGravatar Jason Gross <jgross@mit.edu>2013-11-21 18:39:22 -0500
committerGravatar Pierre Boutillier <pierre.boutillier@ens-lyon.org>2013-12-12 10:09:50 +0100
commitd6c7bf2ec750876c9b35ee9d84840f2ab643dbfe (patch)
tree068e6539f8c9ca5569e2d50a883995d9b6703ed9 /theories/Init/Specif.v
parenta9507f72ffb5cb0b594b097c8d1ed137399fca4a (diff)
Better unification for [projT1] and [proj1_sig].
This way, [fun A (P : A -> Prop) (X : sigT P) => proj1_sig X] unifies with [fun A (P : A -> Prop) (X : sigT P) => projT1 X]. This closes Bug 3043.
Diffstat (limited to 'theories/Init/Specif.v')
-rw-r--r--theories/Init/Specif.v8
1 files changed, 4 insertions, 4 deletions
diff --git a/theories/Init/Specif.v b/theories/Init/Specif.v
index 88c2c264a..ee8111639 100644
--- a/theories/Init/Specif.v
+++ b/theories/Init/Specif.v
@@ -107,11 +107,11 @@ End Projections.
(** [sigT] of a predicate is equivalent to [sig] *)
-Lemma sig_of_sigT : forall (A:Type) (P:A->Prop), sigT P -> sig P.
-Proof. destruct 1 as (x,H); exists x; trivial. Defined.
+Definition sig_of_sigT (A : Type) (P : A -> Prop) (X : sigT P) : sig P
+ := exist P (projT1 X) (projT2 X).
-Lemma sigT_of_sig : forall (A:Type) (P:A->Prop), sig P -> sigT P.
-Proof. destruct 1 as (x,H); exists x; trivial. Defined.
+Definition sigT_of_sig (A : Type) (P : A -> Prop) (X : sig P) : sigT P
+ := existT P (proj1_sig X) (proj2_sig X).
Coercion sigT_of_sig : sig >-> sigT.
Coercion sig_of_sigT : sigT >-> sig.