X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fdemo%2Frealisability.ma;h=93b5e712f24724cee2d004ea2f30e65432e1f16b;hb=3c1ca5620048ad842144fba291f8bc5f0dca7061;hp=a0c021273713bc49a5a05f47f830cc892ca33167;hpb=3deaedbac3407f8b4602b885a6405d4e0cc3e955;p=helm.git diff --git a/helm/software/matita/library/demo/realisability.ma b/helm/software/matita/library/demo/realisability.ma index a0c021273..93b5e712f 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". @@ -106,19 +104,18 @@ theorem correct: ∀F:SP.∀r:type_OF_SP F.modr F r → Prop_OF_SP F. intro; elim F; simplify; [1,2: apply H - | split; + | split; simplify in r H2; [apply H; [ apply (\fst r) | apply (proj1 ? ? H2) ] - | apply H1; + | apply H1;simplify in r H2; [ apply (\snd r) | 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 +128,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