]> matita.cs.unibo.it Git - helm.git/commitdiff
...
authorClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Mon, 5 Oct 2009 14:59:23 +0000 (14:59 +0000)
committerClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Mon, 5 Oct 2009 14:59:23 +0000 (14:59 +0000)
helm/software/matita/nlibrary/sets/partitions.ma
helm/software/matita/nlibrary/sets/sets.ma
helm/software/matita/nlibrary/topology/igft.ma

index 80648f4eacddb68f3207b233654d5d7416db18ce..b8bc9bffe91383868d172d070ed8a8c7f8ee4644 100644 (file)
@@ -19,6 +19,7 @@ include "nat/minus.ma".
 include "datatypes/pairs.ma".
 
 alias symbol "eq" = "setoid eq".
+alias symbol "eq" = "setoid1 eq".
 nrecord partition (A: setoid) : Type[1] ≝ 
  { support: setoid;
    indexes: ext_powerclass support;
@@ -147,6 +148,7 @@ nlemma partition_splits_card:
     nlapply (f_sur ???? (fi nindex) y ?)
      [ alias symbol "refl" = "refl".
 alias symbol "prop1" = "prop11".
+alias symbol "prop2" = "prop21 mem".
 napply (. #‡(†?));##[##2: napply Hni2 |##1: ##skip | nassumption]##]
     *; #nindex2; *; #Hni21; #Hni22;
     nletin xxx ≝ (plus (big_plus (minus n nindex) (λi.λ_.s (S (plus i nindex)))) nindex2);
index 623f676e96d944a4437dee92307063322e33c717..3d03887ec2e4a163292041486a977f596631b40c 100644 (file)
@@ -116,7 +116,8 @@ unification hint 0 ≔ A;
   (* ----------------------------------------------------- *) ⊢  
                  carr1 R ≡ ext_powerclass A.
 
-                
+interpretation "prop21 mem" 'prop2 l r = (prop21 (setoid1_of_setoid ?) (ext_powerclass_setoid ?) ? ? ???? l r).
+      
 (*
 ncoercion ext_carr' : ∀A.∀x:ext_powerclass_setoid A. Ω^A ≝ ext_carr 
 on _x : (carr1 (ext_powerclass_setoid ?)) to (Ω^?). 
@@ -159,7 +160,7 @@ unification hint 0 ≔ A,a,a'
 
 nlemma intersect_is_ext: ∀A. 𝛀^A → 𝛀^A → 𝛀^A.
  #A; #S; #S'; @ (S ∩ S');
- #a; #a'; #Ha; @; *; #H1; #H2; @
+ #a; #a'; #Ha; @; *; #H1; #H2; @ 
   [##1,2: napply (. Ha^-1‡#); nassumption;
 ##|##3,4: napply (. Ha‡#); nassumption]
 nqed.
@@ -192,8 +193,7 @@ unification hint 0 ≔
     fun21 (powerclass_setoid A) (powerclass_setoid A) (powerclass_setoid A) R B C 
   ≡ intersect ? B C.
 
-interpretation "prop21 mem" 'prop2 l r = (prop21 (setoid1_of_setoid ?) (ext_powerclass_setoid ?) ? ???? l r).
-interpretation "prop21 ext" 'prop2 l r = (prop21 (ext_powerclass_setoid ?) (ext_powerclass_setoid ?) ? ???? l r).
+interpretation "prop21 ext" 'prop2 l r = (prop21 (ext_powerclass_setoid ?) (ext_powerclass_setoid ?) ? ? ???? l r).
 
 nlemma intersect_is_ext_morph: 
  ∀A. binary_morphism1 (ext_powerclass_setoid A) (ext_powerclass_setoid A) (ext_powerclass_setoid A).
@@ -231,12 +231,13 @@ unification hint 0 ≔
           
 unification hint 0 ≔
   A, B : CPROP ⊢ iff A B ≡ eq_rel1 ? (eq1 CPROP) A B.    
-*)
     
 nlemma test: ∀U.∀A,B:𝛀^U. A ∩ B = A →
  ∀x,y. x=y → x ∈ A → y ∈ A ∩ B.
  #U; #A; #B; #H; #x; #y; #K; #K2;
-napply (. (prop21 ??? ? ???? K^-1 (H^-1‡#)));
+  alias symbol "prop2" = "prop21 mem".
+  alias symbol "invert" = "setoid1 symmetry".
+  napply (. K^-1‡H);
   nassumption;
 nqed. 
 
@@ -251,14 +252,16 @@ nlemma intersect_ok: ∀A. binary_morphism1 (ext_powerclass_setoid A) (ext_power
       ##|##3,4: napply (. Ha‡#); nassumption]##]
  ##| #a; #a'; #b; #b'; #Ha; #Hb; nwhd; @; #x; nwhd in ⊢ (% → %); #H
       [ alias symbol "invert" = "setoid1 symmetry".
-        napply (. ((#‡Ha^-1)‡(#‡Hb^-1))); nassumption
+        alias symbol "refl" = "refl".
+alias symbol "prop2" = "prop21".
+napply (. ((#‡Ha^-1)‡(#‡Hb^-1))); nassumption
       | napply (. ((#‡Ha)‡(#‡Hb))); nassumption ]##]
 nqed.
 
 (* unfold if intersect, exposing fun21 *)
 alias symbol "hint_decl" = "hint_decl_Type1".
 unification hint 0 ≔ 
-  A : setoid, B,C : qpowerclass A ⊢ 
+  A : setoid, B,C : ext_powerclass A ⊢ 
     pc A (fun21 …
             (mk_binary_morphism1 …
               (λS,S':qpowerclass_setoid A.mk_qpowerclass ? (S ∩ S') (mem_ok' ? (intersect_ok ? S S'))) 
@@ -271,7 +274,6 @@ nlemma test: ∀A:setoid. ∀U,V:qpowerclass A. ∀x,x':setoid1_of_setoid A. x=x
  #A; #U; #V; #x; #x'; #H; #p; napply (. (H^-1‡#)); nassumption.
 nqed.
 *)
-*)
 
 ndefinition image: ∀A,B. (carr A → carr B) → Ω^A → Ω^B ≝
  λA,B:setoid.λf:carr A → carr B.λSa:Ω^A.
index feccae4786bb3ecc5e7ced048b79262209bdda82..8bba1228176ee0eea57c91bb723636f196916928 100644 (file)
@@ -929,6 +929,7 @@ alias symbol "covers" = "new covers".
 alias symbol "covers" = "new covers set".
 alias symbol "covers" = "new covers".
 alias symbol "covers" = "new covers set".
+alias symbol "covers" = "new covers".
 ntheorem new_coverage_infinity:
   ∀A:nAx.∀U:Ω^A.∀a:A. (∃i:𝐈 a. 𝐈𝐦[𝐝 a i] ◃ U) → a ◃ U.
 #A; #U; #a;                                   (** screenshot "n-cov-inf-1". *)  
@@ -1130,7 +1131,8 @@ nelim o;
 ##| #p; #IH; napply (subseteq_intersection_r … IH);
     #x; #Hx; #i; ncases (H … Hx i); #c; *; *; #d; #Ed; #cG;
     @d; napply IH;
-    napply (. Ed^-1‡#); napply cG;    
+    alias symbol "prop2" = "prop21".
+napply (. Ed^-1‡#); napply cG;    
 ##| #a; #i; #f; #Hf; nchange with (G ⊆ { y | ∀x. y ∈ F⎽(f x) }); 
     #b; #Hb; #d; napply (Hf d); napply Hb;
 ##]