]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/nlibrary/topology/cantor.ma
auto is smarter :-)
[helm.git] / helm / software / matita / nlibrary / topology / cantor.ma
index f9655f4ca8f86dbcb36af2a5069679b193310ae9..7864a1ea64b0dda5cf76755251fdaa8114df81eb 100644 (file)
@@ -133,12 +133,10 @@ 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);
-     (* nauto; *) #H1; #H2; #H3; nauto; (* ??? *) ##]
-  napply (replace_char … { x | ϕ a x = O }); ##[##1,2: #x; nrewrite > (H x);
-     (* nauto; *) #E; napply E; ##]
+  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); #x; nauto; ##] #H2;
+ncut (a ◃ ∅); ##[ napply (transitivity … H1); nauto; ##] #H2;
 nlapply (col2_4 …H2); #H3;
 ncut (a ∈ 𝐂 a one); ##[
   nnormalize; ncases (Ph a); nrewrite > (H a); nauto; ##] #H4;