X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=inline;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fdemo%2Frealisability.ma;h=66a96148eadc4cb6c193f0da845ca7ec952b9071;hb=4dc47c9675ffd5fa50296ffaa9b5997501518c98;hp=0cd58c63dce78b1c70fe0640541c2d405c73621e;hpb=27bfea6a5157027d01f57b12ecacdd507b569864;p=helm.git diff --git a/helm/software/matita/library/demo/realisability.ma b/helm/software/matita/library/demo/realisability.ma index 0cd58c63d..66a96148e 100644 --- a/helm/software/matita/library/demo/realisability.ma +++ b/helm/software/matita/library/demo/realisability.ma @@ -12,8 +12,6 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/demo/realisability/". - include "logic/connectives.ma". include "datatypes/constructors.ma". @@ -59,15 +57,13 @@ definition pi2 ≝ match s return λs.P (pi1 ? ? s) with [ sigma_intro _ p ⇒ p]. -notation "hvbox(Σ ident i opt (: ty) break . p)" +notation "hvbox(\Sigma ident i opt (: ty) break . p)" right associative with precedence 20 for @{ 'sigma ${default @{\lambda ${ident i} : $ty. $p} @{\lambda ${ident i} . $p}}}. -(*CSC: the URI must disappear: there is a bug now *) -interpretation "Sigma" 'sigma \eta.x = - (cic:/matita/demo/realisability/sigma.ind#xpointer(1/1) _ x). +interpretation "Sigma" 'sigma \eta.x = (sigma _ x). let rec type_OF_SP F ≝ match F return λF.Type with @@ -116,9 +112,8 @@ theorem correct: ∀F:SP.∀r:type_OF_SP F.modr F r → Prop_OF_SP F. | apply (proj2 ? ? H2) ] ] - | generalize in match H2; clear H2; - change in r with (type_OF_SP s + type_OF_SP s1); - elim r; simplify in H2; + | change in r with (type_OF_SP s + type_OF_SP s1); + elim r in H2 ⊢ %; simplify in H2; [ left; apply H; assumption | right; apply H1; assumption ] @@ -131,8 +126,7 @@ theorem correct: ∀F:SP.∀r:type_OF_SP F.modr F r → Prop_OF_SP F. | skip ] | simplify in r; - generalize in match H1; clear H1; - elim r; + elim r in H1 ⊢ %; apply (ex_intro ? ? a); apply H; assumption