]> matita.cs.unibo.it Git - helm.git/commitdiff
End of curryfication of binary_morphisms.
authorClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Wed, 3 Feb 2010 23:10:57 +0000 (23:10 +0000)
committerClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Wed, 3 Feb 2010 23:10:57 +0000 (23:10 +0000)
helm/software/matita/nlibrary/topology/cantor.ma

index 3e9214f37ce61fc5b9f18086aff74da37a8c91b1..fd300462e918aeecc8add2b9b358a81614a278cf 100644 (file)
@@ -80,19 +80,18 @@ unification hint 0 ≔ ;
   (* -------------- *) ⊢
          S x ≡ A. 
 
-
 ntheorem col2_4 :
-  ∀A:uAx.∀a:A. a ◃ ∅ → ¬ a ∈ 𝐂 a 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); //;
-##] 
+##| #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; //; 
+#A; #a; #U; #V; #HUV; #H; nelim H; /3/; 
 nqed.
 
 ntheorem th3_1: ¬∃a:axs.Z ⊆ S a ∧ S a ⊆ Z.
@@ -103,7 +102,7 @@ ncut (a ◃ Z); ##[
 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;
-//;
+/2/;
 nqed.
 
 include "nat/nat.ma".
@@ -139,20 +138,18 @@ 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; //;
+#A; #U; #V;  #UV; #VU; #a; #aU; nelim aU; /3/;
 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 | 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); //; ##] #H4;
-//;
-nqed.
-
-
+  nnormalize; ncases (Ph a); nrewrite > (H a); /2/; ##] #H4;
+/2/;
+nqed.
\ No newline at end of file