X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Fsets%2Fsets.ma;h=3cf65c11890e6b1d342f004b966bb0166579d33d;hb=dc7c02d8d8678d250a99dd6d012adcd69da63b75;hp=f230957a765e1f877af40378ccd1e9269ac38dd6;hpb=19ae6eb1f1e41a730ce84d47128b0ab0348dae5d;p=helm.git diff --git a/helm/software/matita/nlibrary/sets/sets.ma b/helm/software/matita/nlibrary/sets/sets.ma index f230957a7..3cf65c118 100644 --- a/helm/software/matita/nlibrary/sets/sets.ma +++ b/helm/software/matita/nlibrary/sets/sets.ma @@ -267,6 +267,7 @@ nlemma test: ∀A:setoid. ∀U,V:qpowerclass A. ∀x,x':setoid1_of_setoid A. x=x #A; #U; #V; #x; #x'; #H; #p; napply (. (H^-1‡#)); nassumption. nqed. *) +*) ndefinition image: ∀A,B. (carr A → carr B) → Ω^A → Ω^B ≝ λA,B:setoid.λf:carr A → carr B.λSa:Ω^A. @@ -317,11 +318,11 @@ nlemma first_omomorphism_theorem_functions1: nqed. ndefinition surjective ≝ - λA,B.λS: qpowerclass A.λT: qpowerclass B.λf:unary_morphism A B. + λA,B.λS: ext_powerclass A.λT: ext_powerclass B.λf:unary_morphism A B. ∀y. y ∈ T → ∃x. x ∈ S ∧ f x = y. ndefinition injective ≝ - λA,B.λS: qpowerclass A.λf:unary_morphism A B. + λA,B.λS: ext_powerclass A.λf:unary_morphism A B. ∀x,x'. x ∈ S → x' ∈ S → f x = f x' → x = x'. nlemma first_omomorphism_theorem_functions2: @@ -337,7 +338,7 @@ nlemma first_omomorphism_theorem_functions3: #A; #B; #f; nwhd; #x; #x'; #Hx; #Hx'; #K; nassumption. nqed. -nrecord isomorphism (A, B : setoid) (S: qpowerclass A) (T: qpowerclass B) : Type[0] ≝ +nrecord isomorphism (A, B : setoid) (S: ext_powerclass A) (T: ext_powerclass 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; @@ -365,4 +366,3 @@ ncheck (λA:?. ; }. *) -*) \ No newline at end of file