X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Fsets%2Fsets.ma;h=4e7418fd93bcc17f4777d3bb30a293d674bc43c3;hb=6fe06927f3293bfce4a01a587abd9913e711da88;hp=c8a19354cb74ad72f24a93403bb0ffab5cdb51f6;hpb=021fd04f3bf776eeb4b8da45b3c7103ebabd789d;p=helm.git diff --git a/helm/software/matita/nlibrary/sets/sets.ma b/helm/software/matita/nlibrary/sets/sets.ma index c8a19354c..4e7418fd9 100644 --- a/helm/software/matita/nlibrary/sets/sets.ma +++ b/helm/software/matita/nlibrary/sets/sets.ma @@ -228,9 +228,31 @@ nlemma first_omomorphism_theorem_functions3: #A; #B; #f; nwhd; #x; #x'; #Hx; #Hx'; #K; nassumption. nqed. -nrecord isomorphism (A) (B) (S: qpowerclass A) (T: qpowerclass B) : CProp[0] ≝ +nrecord isomorphism (A, B : setoid) (S: qpowerclass A) (T: qpowerclass B) : Type[0] ≝ { iso_f:> unary_morphism A B; f_closed: ∀x. x ∈ S → iso_f x ∈ T; f_sur: surjective … S T iso_f; f_inj: injective … S iso_f }. + +(* +nrecord isomorphism (A, B : setoid) (S: qpowerclass A) (T: qpowerclass B) : CProp[0] ≝ + { iso_f:> unary_morphism A B; + f_closed: ∀x. x ∈ pc A S → fun1 ?? iso_f x ∈ pc B T}. + + +ncheck (λA:?. + λB:?. + λS:?. + λT:?. + λxxx:isomorphism A B S T. + match xxx + return λxxx:isomorphism A B S T. + ∀x: carr A. + ∀x_72: mem (carr A) (pc A S) x. + mem (carr B) (pc B T) (fun1 A B ((λ_.?) A B S T xxx) x) + with [ mk_isomorphism _ yyy ⇒ yyy ] ). + + ; + }. +*)