]> matita.cs.unibo.it Git - helm.git/commitdiff
Some more lemmas from CerCo.
authorClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Wed, 14 Dec 2011 14:39:09 +0000 (14:39 +0000)
committerClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Wed, 14 Dec 2011 14:39:09 +0000 (14:39 +0000)
matita/matita/lib/basics/types.ma

index d093202b1c340cfe537b06ed9530c32d1dd76c0d..46cc5dd4be8c0c1a1d25374fa7c2fcdc972d2369 100644 (file)
@@ -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] ≝ {