X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fbasics%2Ftypes.ma;h=efc213f95da759cb305ef3d1522ee13691de78e0;hb=dbc57c92512c04b3fd88f8289bb8dbe99b2f90e0;hp=1d38830481c02fa51a4121918f8941fbeaee8285;hpb=dc35589f8448d6a7bd803992e918acc7b720e7fb;p=helm.git diff --git a/matita/matita/lib/basics/types.ma b/matita/matita/lib/basics/types.ma index 1d3883048..efc213f95 100644 --- a/matita/matita/lib/basics/types.ma +++ b/matita/matita/lib/basics/types.ma @@ -9,6 +9,7 @@ \ / GNU General Public License Version 2 V_______________________________________________________________ *) +include "basics/core_notation/pair_2.ma". include "basics/logic.ma". (* void *) @@ -57,20 +58,17 @@ 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 }. -interpretation "DPair" 'sigma x = (DPair ? x). +interpretation "DPair" 'dpair 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). +interpretation "mk_DPair" 'mk_DPair x y = (mk_DPair ?? x y). (* sigma *) record Sig (A:Type[0]) (f:A→Prop) : Type[0] ≝ { - pi1: A + pi1: A (* not a coercion due to problems with Cerco *) ; pi2: f pi1 }. @@ -78,6 +76,13 @@ 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. + +lemma inj_mk_Sig: ∀A,P.∀x. x = mk_Sig A P (pi1 A P x) (pi2 A P x). +#A #P #x cases x // +qed-. (* Prod *) record Prod (A,B:Type[0]) : Type[0] ≝ { @@ -144,10 +149,10 @@ for @{ match $t return λx.x = $t → ? with [ mk_Prod ${fresh xy} ${ident z} match ${fresh xy} return λx. ? = $t → ? with [ mk_Prod ${ident x} ${ident y} ⇒ λ${ident E}.$s ] ] (refl ? $t) }. -notation < "hvbox('let' \nbsp hvbox(〈ident x,ident y,ident z〉 \nbsp'as'\nbsp ident E\nbsp ≝ break t \nbsp 'in' \nbsp) break s)" +notation < "hvbox('let' \nbsp hvbox(〈ident x,ident y,ident z〉 \nbsp 'as' \nbsp ident E\nbsp ≝ break t \nbsp 'in' \nbsp) break s)" with precedence 10 -for @{ match $t return λ${ident x}.$eq $T $x $t → $U with [ mk_Prod (${fresh xy}:$V) (${ident z}:$Z) ⇒ - match ${fresh xy} return λ${ident y}. $eq $R $r $t → ? with [ mk_Prod (${ident x}:$L) (${ident y}:$I) ⇒ +for @{ match $t return λ${ident k}:$X.$eq $T $k $t → $U with [ mk_Prod (${ident xy}:$V) (${ident z}:$Z) ⇒ + match $xy return λ${ident a}. $eq $R $r $t → ? with [ mk_Prod (${ident x}:$L) (${ident y}:$I) ⇒ λ${ident E}:$J.$s ] ] ($refl $A $t) }. notation > "hvbox('let' 〈ident w,ident x,ident y,ident z〉 ≝ t 'in' s)" @@ -213,3 +218,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 ??)).