X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fng_assembly%2Fcommon%2Ftheory.ma;h=2ce5c7036deb6f45b405573846d9f53f48ff418e;hb=34e2c8f59dd7924e15a7746644182d12ad09fed3;hp=093c85e130b850525520cf97a6787be54b988415;hpb=417792b30223b5edd4a9194193c7f34514bd0fa3;p=helm.git diff --git a/helm/software/matita/contribs/ng_assembly/common/theory.ma b/helm/software/matita/contribs/ng_assembly/common/theory.ma index 093c85e13..2ce5c7036 100644 --- a/helm/software/matita/contribs/ng_assembly/common/theory.ma +++ b/helm/software/matita/contribs/ng_assembly/common/theory.ma @@ -525,3 +525,35 @@ ndefinition symmetricT: ∀A,T:Type.∀R:relationT A T.Prop ≝ ndefinition associative : ∀A:Type.∀R:relationT A A.Prop ≝ λA.λR.∀x,y,z:A.R (R x y) z = R x (R y z). + +(* aggiunta per bypassare i punti in cui le dimostrazioni sono equivalenti *) + +ninductive peqv (A:Prop) (x:A) : A → Prop ≝ + prefl_eqv : peqv A x x. + +interpretation "prop equivalence" 'preqv t x y = (peqv t x y). + +(* \equiv *) +notation > "hvbox(a break ≡ b)" + non associative with precedence 45 +for @{ 'preqv ? $a $b }. + +nlemma symmetric_peqv: ∀A:Prop. symmetric A (peqv A). + #A; + nnormalize; + #x; #y; #H; + napply (peqv_ind A x (λ_.?) ? y H); + napply prefl_eqv. +nqed. + +nlemma peqv_ind_r: ∀A:Prop.∀x:A.∀P:A → Prop.P x → ∀y:A.y ≡ x → P y. + #A; #x; #P; #H; #y; #H1; + napply (peqv_ind A x (λ_.?) H y (symmetric_peqv … H1)). +nqed. + +naxiom peqv_ax : ∀P:Prop.∀Q,R:P.Q ≡ R. + +(* uso P x → P y, H e' P x + nrewrite > cioe' napply (peqv_ind ? x (λ_.?) H y (dimostrazione di x ≡ y)); + nrewrite < cioe' napply (peqv_ind_r ? x ? H y (dimostrazione y ≡ x))); +*)