X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Fsets%2Fsets.ma;h=eb99ee4c5c22595075b5dae056155cacb7d856ca;hb=96881c08dcd617524621fb2f241fe38da81f2083;hp=5ba0b08a6f856a40e3199c5cbda4fe439b53def0;hpb=4ae18461e6dfbf0011c062ab56fe85be00f011ec;p=helm.git diff --git a/helm/software/matita/nlibrary/sets/sets.ma b/helm/software/matita/nlibrary/sets/sets.ma index 5ba0b08a6..eb99ee4c5 100644 --- a/helm/software/matita/nlibrary/sets/sets.ma +++ b/helm/software/matita/nlibrary/sets/sets.ma @@ -12,11 +12,9 @@ (* *) (**************************************************************************) -include "logic/connectives.ma". +include "logic/equality.ma". nrecord powerset (A: Type) : Type[1] ≝ { mem: A → CProp }. -(* This is a projection! *) -ndefinition mem ≝ λA.λr:powerset A.match r with [mk_powerset f ⇒ f]. interpretation "powerset" 'powerset A = (powerset A). @@ -34,7 +32,7 @@ nqed. ntheorem subseteq_trans: ∀A.∀S1,S2,S3: Ω \sup A. S1 ⊆ S2 → S2 ⊆ S3 → S1 ⊆ S3. #A; #S1; #S2; #S3; #H12; #H23; #x; #H; - napply (H23 ??); napply (H12 ??); nassumption; + napply H23; napply H12; nassumption; nqed. ndefinition overlaps ≝ λA.λU,V:Ω \sup A.∃x:A.x ∈ U ∧ x ∈ V. @@ -49,8 +47,6 @@ ndefinition union ≝ λA.λU,V:Ω \sup A. {x | x ∈ U ∨ x ∈ V }. interpretation "union" 'union U V = (union ? U V). -(* ndefinition singleton ≝ λA.λa:A.{b | a=b}. interpretation "singleton" 'singl a = (singleton ? a). -*) \ No newline at end of file