]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/nlibrary/topology/cantor.ma
some more work
[helm.git] / helm / software / matita / nlibrary / topology / cantor.ma
index 822c88fc0dfe4f8488c873c9e6130062b262cd7b..b8f8a5543ad57e8cf24ea8ee2f169e63ac94c0aa 100644 (file)
@@ -127,109 +127,31 @@ unification hint 0 ≔ ;
 
 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
-
-
-
-
-
-
-
-
-
-
-
-
-
-
-   
-                     
-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;
-##]
-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;
-    ##]
-##] 
+alias symbol "eq" = "leibnitz's equality".
+alias symbol "eq" = "setoid1 eq".
+alias symbol "covers" = "covers".
+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;  #a; #H1; #H2; #E; nelim E;
+##[ #b; #Hb; @; nauto;
+##| #b; #i; #H3; #H4; @2 i; #c; #Hc; nauto; ##]
 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);
-##]
+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);
+     (* nauto; *) #H1; #H2; #H3; nauto; (* ??? *) ##]
+  napply (replace_char … { x | ϕ a x = O }); ##[##1,2: #x; nrewrite > (H x);
+     (* nauto; *) #E; napply E; ##]
+  napply (axiom_cond … a one); ##] #H1;
+ncut (a ◃ ∅); ##[ napply (transitivity … H1); #x; nauto; ##] #H2;
+nlapply (col2_4 …H2); #H3;
+ncut (a ∈ 𝐂 a one); ##[
+  nnormalize; ncases (Ph a); nrewrite > (H a); nauto; ##] #H4;
+nauto;
 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