]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/nlibrary/topology/cantor.ma
// in place of nauto everywhere
[helm.git] / helm / software / matita / nlibrary / topology / cantor.ma
index 95cf6b3aaf4007d7cc3d096afa43f7ea4a6976d1..e7f9a5a15bed6ccbf0709169558c16a2497ff0dc 100644 (file)
 
 include "topology/igft.ma".
 
+ntheorem axiom_cond: âˆ€A:Ax.∀a:A.∀i:𝐈 a.a â—ƒ đ‚ a i.
+#A; #a; #i; @2 i; #x; #H; @; napply H;
+nqed.
+
+nlemma hint_auto1 : âˆ€A,U,V. (∀x.x âˆˆ U â†’ x â—ƒ V) â†’ cover_set cover A U V.
+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; /3/;
+nqed.
+
+ndefinition emptyset: âˆ€A.Ί^A â‰ ÎťA.{x | False}.
+
+notation "∅" non associative with precedence 90 for @{ 'empty }.
+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/;
+##] 
+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).
+
 ninductive unit : Type[0] â‰ one : unit.
 
-naxiom E: setoid.
-naxiom R: E â†’ ÎŠ^E.
+nrecord uAx : Type[1] â‰ {
+  uax_ : Ax;
+  with_ : âˆ€a:uax_.𝐈 a âˆź unit
+}.
 
-ndefinition axs: Ax.
-@ E (Îť_.unit) (Îťa,x.R a);
+ndefinition uax : uAx â†’ Ax.
+#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.
+
+naxiom A: Type[0].
+naxiom S: A â†’ ÎŠ^A.
+
+ndefinition axs: uAx.
+@; ##[ @ A (Îť_.unit) (Îťa,x.S a); ##| #_; @; ##]
+nqed.
+alias id "S" = "cic:/matita/ng/topology/igft/S.fix(0,0,1)".
 unification hint 0 â‰” ;
          x â‰Ÿ axs  
   (* -------------- *) âŠ˘
-         S x â‰Ą E. 
-         
-ndefinition emptyset: ÎŠ^axs â‰ {x | False}.
+         S x â‰Ą A. 
 
-ndefinition Z: ÎŠ^axs â‰ {x | x â—ƒ emptyset}.
 
+ntheorem col2_4 :
+  âˆ€A:uAx.∀a:A. a â—ƒ âˆ… â†’ ÂŹ a âˆˆ đ‚ a one. 
+#A; #a; #H; nelim H;
+##[ #n; *;
+##| #b; #i_star; #IH1; #IH2; #H3; nlapply (IH2 â€Ś H3); //;
+##] 
+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; //; 
+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);
+  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;
+//;
+nqed.
+
+include "nat/nat.ma".
+
+naxiom phi : nat â†’ nat â†’ nat.
+
+notation > "ϕ" non associative with precedence 90 for @{ 'phi }.
+interpretation "phi" 'phi = phi.
+notation < "ϕ a i" non associative with precedence 90 for @{ 'phi2 $a $i}.
+interpretation "phi2" 'phi2 a i = (phi a i). 
+notation < "ϕ a" non associative with precedence 90 for @{ 'phi1 $a }.
+interpretation "phi2" 'phi1 a = (phi a). 
+
+ndefinition caxs : uAx. 
+@; ##[ @ nat (Îť_.unit); #a; #_; napply { x | Ď• a x = O } ##| #_; @; ##]
+nqed.
+
+
+alias id "S" = "cic:/matita/ng/topology/igft/S.fix(0,0,1)".
+unification hint 0 â‰” ;
+         x â‰Ÿ caxs  
+  (* -------------- *) âŠ˘
+         S x â‰Ą nat. 
+
+naxiom h : nat â†’ nat. 
+
+alias symbol "eq" = "leibnitz's equality".
+alias symbol "eq" = "setoid1 eq".
 alias symbol "covers" = "covers".
-alias symbol "covers" = "covers set".
-alias symbol "covers" = "covers".
-alias symbol "covers" = "covers set".
-ntheorem cover_trans: âˆ€A:Ax.∀a:A.∀U,V. a â—ƒ U â†’ U â—ƒ V â†’ a â—ƒ V.
-#A; #a; #U; #V; #aU; #UV;
-nelim aU;
-##[ #c; #H; napply (UV â€Ś H);
-##| #c; #i; #HCU; #H; napply (cinfinity â€Ś i); napply H;
-##]
-nqed.
-
-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 â—ƒ emptyset); ##[
-  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 *)
+alias symbol "eq" = "leibnitz's equality".
+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; //;
 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); //; ##]
+  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); //; ##] #H4;
+//;
+nqed.
+
+