From: Claudio Sacerdoti Coen Date: Wed, 14 Dec 2011 14:39:09 +0000 (+0000) Subject: Some more lemmas from CerCo. X-Git-Tag: make_still_working~2005 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=910b8e95135bfabb852c2e6d918ff50bdcb98d0a;p=helm.git Some more lemmas from CerCo. --- diff --git a/matita/matita/lib/basics/types.ma b/matita/matita/lib/basics/types.ma index d093202b1..46cc5dd4b 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] ≝ {