]> matita.cs.unibo.it Git - helm.git/commitdiff
...
authorEnrico Tassi <enrico.tassi@inria.fr>
Wed, 21 Oct 2009 14:49:00 +0000 (14:49 +0000)
committerEnrico Tassi <enrico.tassi@inria.fr>
Wed, 21 Oct 2009 14:49:00 +0000 (14:49 +0000)
helm/software/matita/nlibrary/topology/cantor.ma

index 7864a1ea64b0dda5cf76755251fdaa8114df81eb..bd8a04a22804b12589685ceb9597e2159c0c8885 100644 (file)
@@ -87,8 +87,8 @@ 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; nauto; ##] #H1;
 ncut (¬ a ∈ S a); ##[ napply (col2_4 … H1); ##] #H2;
 ncut (a ∈ S a); ##[ napply ZSa; napply H1; ##] #H3;
 nauto;