X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Ftopology%2Fcantor.ma;h=d3dccb8ee07e66248e17e69f21536cc6fb3013fe;hb=e7264f953fcf2cd2bef9057d83add08996d2ce75;hp=3e9214f37ce61fc5b9f18086aff74da37a8c91b1;hpb=aaa04b3cfa6fc3410c953f21c53796f82bb22411;p=helm.git diff --git a/helm/software/matita/nlibrary/topology/cantor.ma b/helm/software/matita/nlibrary/topology/cantor.ma index 3e9214f37..d3dccb8ee 100644 --- a/helm/software/matita/nlibrary/topology/cantor.ma +++ b/helm/software/matita/nlibrary/topology/cantor.ma @@ -41,7 +41,7 @@ ntheorem th2_3 : ∀A:Ax.∀a:A. a ◃ ∅ → ∃i. ¬ a ∈ 𝐂 a i. #A; #a; #H; nelim H; ##[ #n; *; -##| #b; #i_star; #IH1; #IH2; ncases (EM … b i_star); /2/; +##| #b; #i_star; #IH1; #IH2; ncases (EM … b i_star); /3/; ##] nqed. @@ -80,19 +80,18 @@ unification hint 0 ≔ ; (* -------------- *) ⊢ S x ≡ A. - ntheorem col2_4 : - ∀A:uAx.∀a:A. a ◃ ∅ → ¬ a ∈ 𝐂 a one. + ∀A:uAx.∀a:uax A. a ◃ ∅ → ¬ a ∈ 𝐂 a one. #A; #a; #H; nelim H; ##[ #n; *; -##| #b; #i_star; #IH1; #IH2; #H3; nlapply (IH2 … H3); //; -##] +##| #b; #i_star; #IH1; #IH2; #H3; nlapply (IH2 … H3); /2/; +##] nqed. ndefinition Z : Ω^axs ≝ { x | x ◃ ∅ }. ntheorem cover_monotone: ∀A:Ax.∀a:A.∀U,V.U ⊆ V → a ◃ U → a ◃ V. -#A; #a; #U; #V; #HUV; #H; nelim H; //; +#A; #a; #U; #V; #HUV; #H; nelim H; /3/; nqed. ntheorem th3_1: ¬∃a:axs.Z ⊆ S a ∧ S a ⊆ Z. @@ -103,7 +102,7 @@ ncut (a ◃ Z); ##[ ncut (a ◃ ∅); ##[ napply (transitivity … H); nwhd in match Z; //; ##] #H1; ncut (¬ a ∈ S a); ##[ napply (col2_4 … H1); ##] #H2; ncut (a ∈ S a); ##[ napply ZSa; napply H1; ##] #H3; -//; +/2/; nqed. include "nat/nat.ma". @@ -139,20 +138,18 @@ naxiom Ph : ∀x.h x = O \liff x ◃ ∅. nlemma replace_char: ∀A:Ax.∀U,V.U ⊆ V → V ⊆ U → ∀a:A.a ◃ U → a ◃ V. -#A; #U; #V; #UV; #VU; #a; #aU; nelim aU; //; +#A; #U; #V; #UV; #VU; #a; #aU; nelim aU; /3/; nqed. ntheorem th_ch3: ¬∃a:caxs.∀x.ϕ a x = h x. *; #a; #H; ncut (a ◃ { x | x ◃ ∅}); ##[ - napply (replace_char … { x | h x = O }); ##[ ##1,2: #x; ncases (Ph x); //; ##] + napply (replace_char … { x | h x = O }); ##[ ##1,2: #x; ncases (Ph x); /2/; ##] napply (replace_char … { x | ϕ a x = O }); ##[##1,2: #x; nrewrite > (H x); //; ##] napply (axiom_cond … a one); ##] #H1; ncut (a ◃ ∅); ##[ napply (transitivity … H1); //; ##] #H2; nlapply (col2_4 …H2); #H3; ncut (a ∈ 𝐂 a one); ##[ - nnormalize; ncases (Ph a); nrewrite > (H a); //; ##] #H4; -//; -nqed. - - + nnormalize; ncases (Ph a); nrewrite > (H a); /2/; ##] #H4; +/2/; +nqed. \ No newline at end of file