X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Ftopology%2Fcantor.ma;h=6a79dd2d808df2478f6a4ad25568e41e1eb68894;hb=e880d6eab5e1700f4a625ddcd7d0fa8f0cce2dcc;hp=822c88fc0dfe4f8488c873c9e6130062b262cd7b;hpb=e008452eb6b63f53b4eafc13853f7521d411dd00;p=helm.git diff --git a/helm/software/matita/nlibrary/topology/cantor.ma b/helm/software/matita/nlibrary/topology/cantor.ma index 822c88fc0..6a79dd2d8 100644 --- a/helm/software/matita/nlibrary/topology/cantor.ma +++ b/helm/software/matita/nlibrary/topology/cantor.ma @@ -1,4 +1,16 @@ - +(**************************************************************************) +(* ___ *) +(* ||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". @@ -7,18 +19,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; nauto. +nnormalize; /2/. 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; -##[ #c; #H; nauto; -##| #c; #i; #HCU; #H; @2 i; nauto; -##] +#A; #a; #U; #V; #aU; #UV; nelim aU; /3/; nqed. ndefinition emptyset: ∀A.Ω^A ≝ λA.{x | False}. @@ -28,15 +36,12 @@ interpretation "empty" 'empty = (emptyset ?). naxiom EM : ∀A:Ax.∀a:A.∀i_star.(a ∈ 𝐂 a i_star) ∨ ¬( a ∈ 𝐂 a i_star). +alias symbol "covers" = "covers". 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: (* nauto; *) #W; @i_star; napply W; - ##| nauto; - ##] +##| #b; #i_star; #IH1; #IH2; ncases (EM … b i_star); /3/; ##] nqed. @@ -56,7 +61,8 @@ nrecord uAx : Type[1] ≝ { }. ndefinition uax : uAx → Ax. -#A; @ (uax_ A) (λx.unit); #a; #_; napply (𝐂 a ?); nlapply one; ncases (with_ A a); nauto; +#A; @ (uax_ A) (λx.unit); #a; #_; +napply (𝐂 a ?); nlapply one; ncases (with_ A a); //; nqed. ncoercion uax : ∀u:uAx. Ax ≝ uax on _u : uAx to Ax. @@ -72,34 +78,32 @@ alias id "S" = "cic:/matita/ng/topology/igft/S.fix(0,0,1)". unification hint 0 ≔ ; x ≟ axs (* -------------- *) ⊢ - S x ≡ A. - + S (uax x) ≡ A. (* XXX: bug coercions/ disamb multipasso che ne fa 1 solo*) ntheorem col2_4 : - ∀A:uAx.∀a:A. a ◃ ∅ → ¬ a ∈ 𝐂 a ?. ##[ (* bug *) ##2: nnormalize; napply 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); #H4; nauto; -##] +##| #b; #i_star; #IH1; #IH2; #H3; nlapply (IH2 … H3); /2/; +##] nqed. -ndefinition Z : Ω^axs ≝ { x | x ◃ ∅ }. +(* bug interpretazione non aggiunta per ∅ *) +ndefinition Z : Ω^axs ≝ { x | x ◃ (emptyset ?) }. ntheorem cover_monotone: ∀A:Ax.∀a:A.∀U,V.U ⊆ V → a ◃ U → a ◃ V. -#A; #a; #U; #V; #HUV; #H; nelim H; -##[ nauto; -##| #b; #i; #HCU; #W; @2 i; #x; nauto; ##] +#A; #a; #U; #V; #HUV; #H; nelim H; /3/; nqed. ntheorem th3_1: ¬∃a:axs.Z ⊆ S a ∧ S a ⊆ Z. *; #a; *; #ZSa; #SaZ; ncut (a ◃ Z); ##[ nlapply (axiom_cond … a one); #AxCon; nchange in AxCon with (a ◃ S a); - (* nauto; *) napply (cover_monotone … AxCon); nassumption; ##] #H; -ncut (a ◃ ∅); ##[ napply (transitivity … H); #x; #E; napply E; ##] #H1; + napply (cover_monotone … AxCon); nassumption; ##] #H; +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; -nauto; +/2/; nqed. include "nat/nat.ma". @@ -123,113 +127,30 @@ alias id "S" = "cic:/matita/ng/topology/igft/S.fix(0,0,1)". unification hint 0 ≔ ; x ≟ caxs (* -------------- *) ⊢ - S x ≡ nat. + S (uax x) ≡ nat. naxiom h : nat → nat. -naxiom Ph : ∀x.h x = O → x ◃ ∅. - -ninductive eq2 (A : Type[1]) (a : A) : A → CProp[0] ≝ -| refl2 : eq2 A a a. - -interpretation "eq2" 'eq T a b = (eq2 T a b). - -ntheorem th_ch3: ¬∃a:caxs.∀x.ϕ a x = h x. -*; #a; #H; -ncut ((𝐂 a one) ⊆ { x | x ◃ ∅ }); (* bug *) -nchange in xx with { x | h x - - - - - - +alias symbol "eq" = "leibnitz's equality". +alias symbol "eq" = "setoid1 eq". +alias symbol "covers" = "covers". +alias symbol "eq" = "leibnitz's equality". +naxiom Ph : ∀x.h x = O \liff x ◃ ∅. - - - - - - - - - -ntheorem cantor: ∀a:axs. ¬ (Z ⊆ R a ∧ R a ⊆ Z). -#a; *; #ZRa; #RaZ; -ncut (a ◃ R a); ##[ @2; ##[ napply one; ##] #x; #H; @; napply H; ##] #H1; -ncut (a ◃ ∅); ##[ - napply (cover_trans … H1); - #x; #H; nlapply (RaZ … H); #ABS; napply ABS; ##] #H2; -ncut (a ∈ R a); ##[ napply ZRa; napply H2; ##] #H3; -nelim H2 in H3; -##[ nauto. -##| nnormalize; nauto. ##] (* se lo lancio su entrambi fallisce di width *) -nqed. - -ninductive deduct (A : nAx) (U : Ω^A) : A → CProp[0] ≝ -| drefl : ∀a.a ∈ U → deduct A U a -| dcut : ∀a.∀i:𝐈 a. (∀y:𝐃 a i.deduct A U (𝐝 a i y)) → deduct A U a. - -notation " a ⊢ b " non associative with precedence 45 for @{ 'deduct $a $b }. -interpretation "deduct" 'deduct a b = (deduct ? b a). - -ntheorem th2_3_1 : ∀A:nAx.∀a:A.∀i:𝐈 a. a ⊢ 𝐈𝐦[𝐝 a i]. -#A; #a; #i; -ncut (∀y:𝐃 a i.𝐝 a i y ⊢ 𝐈𝐦[𝐝 a i]); ##[ #y; @; @y; @; ##] #H1; -napply (dcut … i); nassumption; -nqed. - -ntheorem th2_3_2 : - ∀A:nAx.∀a:A.∀i:𝐈 a.∀U,V. a ⊢ U → (∀u.u ∈ U → u ⊢ V) → a ⊢ V. -#A; #a; #i; #U; #V; #aU; #xUxV; -nelim aU; -##[ nassumption; -##| #b; #i; #dU; #dV; @2 i; nassumption; -##] +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/; nqed. -ntheorem th2_3 : - ∀A:nAx. - (∀a:A.∀i_star.(∃y:𝐃 a i_star.𝐝 a i_star y = a) ∨ ¬(∃y:𝐃 a i_star.𝐝 a i_star y = a)) → - ∀a:A. a ◃ ∅ → ∃i:𝐈 a. ¬ a \in Z -#A; #EM; #a; #H; nelim H; -##[ #n; *; -##| #b; #i_star; #IH1; #IH2; - ncases (EM b i_star); - ##[##2: #W; @i_star; napply W; - ##| *; #y_star; #E; nlapply (IH2 y_star); nrewrite > E; #OK; napply OK; - ##] -##] -nqed. - -ninductive eq1 (A : Type[0]) : Type[0] → CProp[0] ≝ -| refl1 : eq1 A A. - -notation "hvbox( a break ∼ b)" non associative with precedence 40 -for @{ 'eqT $a $b }. - -interpretation "eq between types" 'eqT a b = (eq1 a b). - -nrecord uAx : Type[1] ≝ { - uax_ : Ax; - with_ : ∀a:uax_.𝐈 a ∼ unit -}. - -ndefinition uax : uAx → Ax. -*; #A; #E; @ A (λx.unit); #a; ncases (E a); -##[ #i; napply (𝐃 a i); -##| #i; nnormalize; #j; napply (𝐝 a i j); -##] -nqed. - -ncoercion uax : ∀u:unAx. nAx ≝ uax on _u : unAx to nAx. - - -nlemma cor_2_5 : ∀A:unAx.∀a:A.∀i.a ⊢ ∅ → ¬(a ∈ 𝐈𝐦[𝐝 a i]). -#A; #a; #i; #H; nelim H in i; -##[ #w; *; -##| #a; #i; #IH1; #IH2; - - - - \ No newline at end of file +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 (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); /2/; ##] #H4; +/2/; +nqed. \ No newline at end of file