X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Ftopology%2Fcantor.ma;h=bd8a04a22804b12589685ceb9597e2159c0c8885;hb=a90c31c1b53222bd6d57360c5ba5c2d0fe7d5207;hp=6a79dd2d808df2478f6a4ad25568e41e1eb68894;hpb=4377e950998c9c63937582952a79975947aa9a45;p=helm.git diff --git a/helm/software/matita/nlibrary/topology/cantor.ma b/helm/software/matita/nlibrary/topology/cantor.ma index 6a79dd2d8..bd8a04a22 100644 --- a/helm/software/matita/nlibrary/topology/cantor.ma +++ b/helm/software/matita/nlibrary/topology/cantor.ma @@ -1,16 +1,4 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) + include "topology/igft.ma". @@ -19,14 +7,14 @@ ntheorem axiom_cond: ∀A:Ax.∀a:A.∀i:𝐈 a.a ◃ 𝐂 a i. nqed. nlemma hint_auto1 : ∀A,U,V. (∀x.x ∈ U → x ◃ V) → cover_set cover A U V. -nnormalize; /2/. +nnormalize; nauto. nqed. alias symbol "covers" (instance 1) = "covers". alias symbol "covers" (instance 2) = "covers set". alias symbol "covers" (instance 3) = "covers". ntheorem transitivity: ∀A:Ax.∀a:A.∀U,V. a ◃ U → U ◃ V → a ◃ V. -#A; #a; #U; #V; #aU; #UV; nelim aU; /3/; +#A; #a; #U; #V; #aU; #UV; nelim aU; nauto depth=4; nqed. ndefinition emptyset: ∀A.Ω^A ≝ λA.{x | False}. @@ -41,7 +29,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); /3/; +##| #b; #i_star; #IH1; #IH2; ncases (EM … b i_star); nauto; ##] nqed. @@ -62,7 +50,7 @@ nrecord uAx : Type[1] ≝ { ndefinition uax : uAx → Ax. #A; @ (uax_ A) (λx.unit); #a; #_; -napply (𝐂 a ?); nlapply one; ncases (with_ A a); //; +napply (𝐂 a ?); nlapply one; ncases (with_ A a); nauto; nqed. ncoercion uax : ∀u:uAx. Ax ≝ uax on _u : uAx to Ax. @@ -78,21 +66,21 @@ alias id "S" = "cic:/matita/ng/topology/igft/S.fix(0,0,1)". unification hint 0 ≔ ; x ≟ axs (* -------------- *) ⊢ - S (uax x) ≡ A. (* XXX: bug coercions/ disamb multipasso che ne fa 1 solo*) + S x ≡ A. + ntheorem col2_4 : - ∀A:uAx.∀a:uax A. a ◃ ∅ → ¬ a ∈ 𝐂 a one. + ∀A:uAx.∀a:A. a ◃ ∅ → ¬ a ∈ 𝐂 a one. #A; #a; #H; nelim H; ##[ #n; *; -##| #b; #i_star; #IH1; #IH2; #H3; nlapply (IH2 … H3); /2/; -##] +##| #b; #i_star; #IH1; #IH2; #H3; nlapply (IH2 … H3); nauto; +##] nqed. -(* bug interpretazione non aggiunta per ∅ *) -ndefinition Z : Ω^axs ≝ { x | x ◃ (emptyset ?) }. +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; /3/; +#A; #a; #U; #V; #HUV; #H; nelim H; nauto depth=4; nqed. ntheorem th3_1: ¬∃a:axs.Z ⊆ S a ∧ S a ⊆ Z. @@ -100,10 +88,10 @@ ntheorem th3_1: ¬∃a:axs.Z ⊆ S a ∧ S a ⊆ Z. ncut (a ◃ Z); ##[ nlapply (axiom_cond … a one); #AxCon; nchange in AxCon with (a ◃ S a); napply (cover_monotone … AxCon); nassumption; ##] #H; -ncut (a ◃ ∅); ##[ napply (transitivity … H); nwhd in match Z; //; ##] #H1; +ncut (a ◃ ∅); ##[ napply (transitivity … H); nwhd in match Z; nauto; ##] #H1; ncut (¬ a ∈ S a); ##[ napply (col2_4 … H1); ##] #H2; ncut (a ∈ S a); ##[ napply ZSa; napply H1; ##] #H3; -/2/; +nauto; nqed. include "nat/nat.ma". @@ -127,7 +115,7 @@ alias id "S" = "cic:/matita/ng/topology/igft/S.fix(0,0,1)". unification hint 0 ≔ ; x ≟ caxs (* -------------- *) ⊢ - S (uax x) ≡ nat. + S x ≡ nat. naxiom h : nat → nat. @@ -139,18 +127,20 @@ 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; /3/; +#A; #U; #V; #UV; #VU; #a; #aU; nelim aU; nauto; 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); /2/; ##] - napply (replace_char … { x | ϕ a x = O }); ##[##1,2: #x; nrewrite > (H x); //; ##] + napply (replace_char … { x | h x = O }); ##[ ##1,2: #x; ncases (Ph x); nauto; ##] + napply (replace_char … { x | ϕ a x = O }); ##[##1,2: #x; nrewrite > (H x); nauto; ##] napply (axiom_cond … a one); ##] #H1; -ncut (a ◃ ∅); ##[ napply (transitivity … H1); //; ##] #H2; +ncut (a ◃ ∅); ##[ napply (transitivity … H1); nauto; ##] #H2; nlapply (col2_4 …H2); #H3; ncut (a ∈ 𝐂 a one); ##[ - nnormalize; ncases (Ph a); nrewrite > (H a); /2/; ##] #H4; -/2/; -nqed. \ No newline at end of file + nnormalize; ncases (Ph a); nrewrite > (H a); nauto; ##] #H4; +nauto; +nqed. + +