X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Ftopology%2Figft.ma;h=59ccd56019a34b690063b71f830dfbed5729b4ea;hb=2f67829c12dea538114bb848e0275d220dba601b;hp=81aefd555f02c320104e47f3d272fbf5301c6673;hpb=6677f4b55d7d949ab7b8f7133f701bcc1593bbeb;p=helm.git diff --git a/helm/software/matita/nlibrary/topology/igft.ma b/helm/software/matita/nlibrary/topology/igft.ma index 81aefd555..59ccd5601 100644 --- a/helm/software/matita/nlibrary/topology/igft.ma +++ b/helm/software/matita/nlibrary/topology/igft.ma @@ -116,26 +116,6 @@ D*) include "sets/sets.ma". -(*HIDE*) -(* move away *) -nlemma subseteq_intersection_l: ∀A.∀U,V,W:Ω^A.U ⊆ W ∨ V ⊆ W → U ∩ V ⊆ W. -#A; #U; #V; #W; *; #H; #x; *; #xU; #xV; napply H; nassumption; -nqed. - -nlemma subseteq_union_l: ∀A.∀U,V,W:Ω^A.U ⊆ W → V ⊆ W → U ∪ V ⊆ W. -#A; #U; #V; #W; #H; #H1; #x; *; #Hx; ##[ napply H; ##| napply H1; ##] nassumption; -nqed. - -nlemma subseteq_intersection_r: ∀A.∀U,V,W:Ω^A.W ⊆ U → W ⊆ V → W ⊆ U ∩ V. -#A; #U; #V; #W; #H1; #H2; #x; #Hx; @; ##[ napply H1; ##| napply H2; ##] nassumption; -nqed. - -ninductive sigma (A : Type[0]) (P : A → CProp[0]) : Type[0] ≝ - sig_intro : ∀x:A.P x → sigma A P. - -interpretation "sigma" 'sigma \eta.p = (sigma ? p). -(*UNHIDE*) - (*D Some basic results that we will use are also part of the sets library: @@ -927,6 +907,9 @@ D*) alias symbol "covers" = "new covers set". alias symbol "covers" = "new covers". alias symbol "covers" = "new covers set". +alias symbol "covers" = "new covers". +alias symbol "covers" = "new covers set". +alias symbol "covers" = "new covers". ntheorem new_coverage_infinity: ∀A:nAx.∀U:Ω^A.∀a:A. (∃i:𝐈 a. 𝐈𝐦[𝐝 a i] ◃ U) → a ◃ U. #A; #U; #a; (** screenshot "n-cov-inf-1". *) @@ -1119,7 +1102,7 @@ before doing the induction over `o`. D*) ntheorem max_new_fished: - ∀A:nAx.∀G:qpowerclass_setoid A.∀F:Ω^A.G ⊆ F → (∀a.a ∈ G → ∀i.𝐈𝐦[𝐝 a i] ≬ G) → G ⊆ ⋉F. + ∀A:nAx.∀G:ext_powerclass_setoid A.∀F:Ω^A.G ⊆ F → (∀a.a ∈ G → ∀i.𝐈𝐦[𝐝 a i] ≬ G) → G ⊆ ⋉F. #A; #G; #F; #GF; #H; #b; #HbG; #o; ngeneralize in match HbG; ngeneralize in match b; nchange with (G ⊆ F⎽o); @@ -1128,7 +1111,8 @@ nelim o; ##| #p; #IH; napply (subseteq_intersection_r … IH); #x; #Hx; #i; ncases (H … Hx i); #c; *; *; #d; #Ed; #cG; @d; napply IH; - napply (. Ed^-1‡#); napply cG; + alias symbol "prop2" = "prop21". +napply (. Ed^-1‡#); napply cG; ##| #a; #i; #f; #Hf; nchange with (G ⊆ { y | ∀x. y ∈ F⎽(f x) }); #b; #Hb; #d; napply (Hf d); napply Hb; ##]