]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/nlibrary/topology/cantor.ma
minimization.ma
[helm.git] / helm / software / matita / nlibrary / topology / cantor.ma
index 822c88fc0dfe4f8488c873c9e6130062b262cd7b..d3dccb8ee07e66248e17e69f21536cc6fb3013fe 100644 (file)
@@ -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.
@@ -74,32 +80,29 @@ unification hint 0 â‰” ;
   (* -------------- *) âŠ˘
          S x â‰Ą A. 
 
-
 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 â—ƒ âˆ… }.
 
 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".
@@ -127,109 +130,26 @@ 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
-
-
-
-
-
-
+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