]> matita.cs.unibo.it Git - helm.git/commitdiff
Dependent pairs (i.e. Sigma types in Type[0]) are back, but are not the default
authorClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Tue, 13 Dec 2011 15:10:12 +0000 (15:10 +0000)
committerClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Tue, 13 Dec 2011 15:10:12 +0000 (15:10 +0000)
ones.

matita/matita/lib/basics/types.ma

index 348bf92de73b52182447045a9efe08b9c590190b..1d38830481c02fa51a4121918f8941fbeaee8285 100644 (file)
@@ -55,6 +55,19 @@ 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" 'sigma x = (DPair ? x).
+
+notation "hvbox(« term 19 a, break term 19 b»)" 
+with precedence 90 for @{ 'dp $a $b }.
+
+interpretation "mk_DPair" 'dp x y = (mk_DPair ?? x y).
+
 (* sigma *)
 record Sig (A:Type[0]) (f:A→Prop) : Type[0] ≝ {
     pi1: A
@@ -63,9 +76,6 @@ 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).
 
 (* Prod *)