]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/nlibrary/topology/cantor.ma
Release 0.5.9.
[helm.git] / helm / software / matita / nlibrary / topology / cantor.ma
index 6a79dd2d808df2478f6a4ad25568e41e1eb68894..bd8a04a22804b12589685ceb9597e2159c0c8885 100644 (file)
@@ -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.
+
+