X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fbasics%2Ftypes.ma;h=1dcd79b7a49892702ba1290f86f3cb0099d3e335;hb=927bda3b4b7fe5f521ae73eb008a746e8606a0b4;hp=d093202b1c340cfe537b06ed9530c32d1dd76c0d;hpb=c9c6cae5121a25b05450ea42578f14f74569cfbf;p=helm.git diff --git a/matita/matita/lib/basics/types.ma b/matita/matita/lib/basics/types.ma index d093202b1..1dcd79b7a 100644 --- a/matita/matita/lib/basics/types.ma +++ b/matita/matita/lib/basics/types.ma @@ -57,7 +57,7 @@ lemma refute_none_by_refl : ∀A,B:Type[0]. ∀P:A → B. ∀Q:B → Type[0]. (* dependent pair *) record DPair (A:Type[0]) (f:A→Type[0]) : Type[0] ≝ { - dpi1: A + dpi1:> A ; dpi2: f dpi1 }. @@ -75,6 +75,10 @@ interpretation "Sigma" 'sigma x = (Sig ? x). interpretation "mk_Sig" 'dp x y = (mk_Sig ?? x y). +lemma sub_pi2 : ∀A.∀P,P':A → Prop. (∀x.P x → P' x) → ∀x:Σx:A.P x. P' (pi1 … x). +#A #P #P' #H1 * #x #H2 @H1 @H2 +qed. + (* Prod *) record Prod (A,B:Type[0]) : Type[0] ≝ { @@ -210,3 +214,10 @@ lemma pair_destruct_2: ∀A,B.∀a:A.∀b:B.∀c. 〈a,b〉 = c → b = \snd c. #A #B #a #b *; /2/ qed. + +lemma coerc_pair_sigma: + ∀A,B,P. ∀p:A × B. P (\snd p) → A × (Σx:B.P x). +#A #B #P * #a #b #p % [@a | /2/] +qed. +coercion coerc_pair_sigma:∀A,B,P. ∀p:A × B. P (\snd p) → A × (Σx:B.P x) +≝ coerc_pair_sigma on p: (? × ?) to (? × (Sig ??)).