]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/basics/types.ma
Some more lemmas from CerCo.
[helm.git] / matita / matita / lib / basics / types.ma
index 348bf92de73b52182447045a9efe08b9c590190b..46cc5dd4be8c0c1a1d25374fa7c2fcdc972d2369 100644 (file)
@@ -55,6 +55,16 @@ lemma refute_none_by_refl : ∀A,B:Type[0]. ∀P:A → B. ∀Q:B → Type[0]. 
 | #a #H #p normalize @p @refl
 ] qed.
 
+(* dependent pair *)
+record DPair (A:Type[0]) (f:A→Type[0]) : Type[0] ≝ {
+    dpi1:> A
+  ; dpi2: f dpi1
+  }.
+
+interpretation "DPair" 'dpair x = (DPair ? x).
+
+interpretation "mk_DPair" 'mk_DPair x y = (mk_DPair ?? x y).
+
 (* sigma *)
 record Sig (A:Type[0]) (f:A→Prop) : Type[0] ≝ {
     pi1: A
@@ -63,11 +73,12 @@ record Sig (A:Type[0]) (f:A→Prop) : Type[0] ≝ {
   
 interpretation "Sigma" 'sigma x = (Sig ? x).
 
-notation "hvbox(« term 19 a, break term 19 b»)" 
-with precedence 90 for @{ 'dp $a $b }.
-
 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] ≝ {