X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fbasics%2Ftypes.ma;h=a5e35c9d218a5d756319dd429dd7e1b6d4e8363e;hb=376fd7774ef0fa2f30a4afb25aab6158e3cd04b7;hp=b7a022ad201056d7d5d1842ea9b3def64d0d862b;hpb=d2b59bd89f761a16a2dbc663f446b4f95c767b83;p=helm.git diff --git a/matita/matita/lib/basics/types.ma b/matita/matita/lib/basics/types.ma index b7a022ad2..a5e35c9d2 100644 --- a/matita/matita/lib/basics/types.ma +++ b/matita/matita/lib/basics/types.ma @@ -67,7 +67,7 @@ interpretation "mk_DPair" 'mk_DPair x y = (mk_DPair ?? x y). (* sigma *) record Sig (A:Type[0]) (f:A→Prop) : Type[0] ≝ { - pi1: A + pi1: A (* not a coercion due to problems with Cerco *) ; pi2: f pi1 }. @@ -79,6 +79,9 @@ lemma sub_pi2 : ∀A.∀P,P':A → Prop. (∀x.P x → P' x) → ∀x:Σx:A.P x. #A #P #P' #H1 * #x #H2 @H1 @H2 qed. +lemma inj_mk_Sig: ∀A,P.∀x. x = mk_Sig A P (pi1 A P x) (pi2 A P x). +#A #P #x cases x // +qed-. (* Prod *) record Prod (A,B:Type[0]) : Type[0] ≝ {