]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/nlibrary/sets/sets.ma
non uniform coercions landed in hints_declaration.ma, setoids and sets library
[helm.git] / helm / software / matita / nlibrary / sets / sets.ma
index 544db6a562bb4cd266d38082ace2ac0dc6197fc0..dcb740921fe0bc323e3559d9feb127abd8eba312 100644 (file)
@@ -62,8 +62,6 @@ ndefinition powerclass_setoid: Type[0] → setoid1.
  #A; @(Ω^A);//.
 nqed.
 
-include "hints_declaration.ma". 
-
 alias symbol "hint_decl" = "hint_decl_Type2".
 unification hint 0 ≔ A;
   R ≟ (mk_setoid1 (Ω^A) (eq1 (powerclass_setoid A)))
@@ -108,10 +106,10 @@ unification hint 0 ≔ A;
                  carr1 R ≡ ext_powerclass A.
 
 nlemma mem_ext_powerclass_setoid_is_morph: 
- ∀A. (setoid1_of_setoid A) ⇒_1 (unary_morphism1_setoid1 (ext_powerclass_setoid A) CPROP).
- #A; napply (mk_binary_morphism1 … (λx:setoid1_of_setoid A.λS: 𝛀^A. x ∈ S));
- #a; #a'; #b; #b'; #Ha; *; #Hb1; #Hb2; @; #H
-  [ napply (. (ext_prop … Ha^-1)) | napply (. (ext_prop … Ha)) ] /2/.
+ ∀A. (setoid1_of_setoid A) ⇒_1 ((𝛀^A) ⇒_1 CPROP).
+#A; napply (mk_binary_morphism1 … (λx:setoid1_of_setoid A.λS: 𝛀^A. x ∈ S));
+#a; #a'; #b; #b'; #Ha; *; #Hb1; #Hb2; @; #H
+[ napply (. (ext_prop … Ha^-1)) | napply (. (ext_prop … Ha)) ] /2/.
 nqed.
 
 unification hint 0 ≔  AA, x, S;  
@@ -135,9 +133,7 @@ nlemma set_ext : ∀S.∀A,B:Ω^S.A =_1 B → ∀x:S.(x ∈ A) =_1 (x ∈ B).
 nlemma ext_set : ∀S.∀A,B:Ω^S.(∀x:S. (x ∈ A) = (x ∈ B)) → A = B.
 #S A B H; @; #x; ncases (H x); #H1 H2; ##[ napply H1 | napply H2] nqed.
 
-nlemma subseteq_is_morph: ∀A. 
- (ext_powerclass_setoid A) ⇒_1 
-   (unary_morphism1_setoid1 (ext_powerclass_setoid A) CPROP).
+nlemma subseteq_is_morph: ∀A.  𝛀^A ⇒_1 𝛀^A ⇒_1 CPROP.
  #A; napply (mk_binary_morphism1 … (λS,S':𝛀^A. S ⊆ S'));
  #a; #a'; #b; #b'; *; #H1; #H2; *; /5/ by mk_iff, sym1, subseteq_trans;
 nqed.
@@ -161,11 +157,10 @@ unification hint 0 ≔
   R ≟ (mk_ext_powerclass ? (B ∩ C) (ext_prop ? (intersect_is_ext ? B C))) 
   (* ------------------------------------------*)  ⊢ 
     ext_carr A R ≡ intersect ? (ext_carr ? B) (ext_carr ? C).
-
-nlemma intersect_is_morph:
- ∀A. (powerclass_setoid A) ⇒_1 (unary_morphism1_setoid1 (powerclass_setoid A) (powerclass_setoid A)).
- #A; napply (mk_binary_morphism1 … (λS,S'. S ∩ S'));
- #a; #a'; #b; #b'; *; #Ha1; #Ha2; *; #Hb1; #Hb2; @; #x; nnormalize; *;/3/.
+    
+nlemma intersect_is_morph: ∀A. Ω^A ⇒_1 Ω^A ⇒_1 Ω^A.
+#A; napply (mk_binary_morphism1 … (λS,S'. S ∩ S'));
+#a; #a'; #b; #b'; *; #Ha1; #Ha2; *; #Hb1; #Hb2; @; #x; nnormalize; *;/3/.
 nqed.
 
 alias symbol "hint_decl" = "hint_decl_Type1".
@@ -180,10 +175,7 @@ interpretation "prop21 ext" 'prop2 l r =
  (prop11 (ext_powerclass_setoid ?)
   (unary_morphism1_setoid1 (ext_powerclass_setoid ?) ?) ? ?? l ?? r).
 
-nlemma intersect_is_ext_morph: 
- ∀A. 
- (ext_powerclass_setoid A) ⇒_1
-  (unary_morphism1_setoid1 (ext_powerclass_setoid A) (ext_powerclass_setoid A)).
+nlemma intersect_is_ext_morph: ∀A. 𝛀^A ⇒_1 𝛀^A ⇒_1 𝛀^A.
  #A; napply (mk_binary_morphism1 … (intersect_is_ext …));
  #a; #a'; #b; #b'; #Ha; #Hb; napply (prop11 … (intersect_is_morph A)); nassumption.
 nqed.
@@ -203,9 +195,7 @@ unification hint 1 ≔
    (* ------------------------------------------------------*) ⊢ 
             ext_carr AA (R B C) ≡ intersect A BB CC.
 
-nlemma union_is_morph : 
- ∀A. (powerclass_setoid A) ⇒_1 (unary_morphism1_setoid1 (powerclass_setoid A) (powerclass_setoid A)).
-(*XXX ∀A.Ω^A ⇒_1 (Ω^A ⇒_1 Ω^A). avec non-unif-coerc*)
+nlemma union_is_morph : ∀A. Ω^A ⇒_1 (Ω^A ⇒_1 Ω^A).
 #X; napply (mk_binary_morphism1 … (λA,B.A ∪ B));
 #A1 A2 B1 B2 EA EB; napply ext_set; #x;
 nchange in match (x ∈ (A1 ∪ B1)) with (?∨?);
@@ -235,10 +225,7 @@ unification hint 0 ≔ S:Type[0], A,B:Ω^S;
 (*--------------------------------------------------------------------------*) ⊢
    fun11 ?? (fun11 ?? MM A) B ≡ A ∪ B.
    
-nlemma union_is_ext_morph:∀A.
- (ext_powerclass_setoid A) ⇒_1
-  (unary_morphism1_setoid1 (ext_powerclass_setoid A) (ext_powerclass_setoid A)).
-(*XXX ∀A:setoid.𝛀^A ⇒_1 (𝛀^A ⇒_1 𝛀^A). with coercion non uniformi *)
+nlemma union_is_ext_morph:∀A.𝛀^A ⇒_1 𝛀^A ⇒_1 𝛀^A.
 #A; napply (mk_binary_morphism1 …  (union_is_ext …));
 #x1 x2 y1 y2 Ex Ey; napply (prop11 … (union_is_morph A)); nassumption.
 nqed.
@@ -386,16 +373,6 @@ nrecord isomorphism (A, B : setoid) (S: ext_powerclass A) (T: ext_powerclass B)
    f_inj: injective … S iso_f
  }.
 
-nlemma subseteq_intersection_l: ∀A.∀U,V,W:Ω^A.U ⊆ W ∨ V ⊆ W → U ∩ V ⊆ W.
-#A; #U; #V; #W; *; #H; #x; *; /2/.
-nqed.
-
-nlemma subseteq_union_l: ∀A.∀U,V,W:Ω^A.U ⊆ W → V ⊆ W → U ∪ V ⊆ W.
-#A; #U; #V; #W; #H; #H1; #x; *; /2/.
-nqed.
-
-nlemma subseteq_intersection_r: ∀A.∀U,V,W:Ω^A.W ⊆ U → W ⊆ V → W ⊆ U ∩ V.
-/3/. nqed.
 
 (*
 nrecord isomorphism (A, B : setoid) (S: qpowerclass A) (T: qpowerclass B) : CProp[0] ≝
@@ -417,4 +394,36 @@ ncheck (λA:?.
    
    ;
  }.
-*)
\ No newline at end of file
+*)
+
+(* Set theory *)
+
+nlemma subseteq_intersection_l: ∀A.∀U,V,W:Ω^A.U ⊆ W ∨ V ⊆ W → U ∩ V ⊆ W.
+#A; #U; #V; #W; *; #H; #x; *; /2/.
+nqed.
+
+nlemma subseteq_union_l: ∀A.∀U,V,W:Ω^A.U ⊆ W → V ⊆ W → U ∪ V ⊆ W.
+#A; #U; #V; #W; #H; #H1; #x; *; /2/.
+nqed.
+
+nlemma subseteq_intersection_r: ∀A.∀U,V,W:Ω^A.W ⊆ U → W ⊆ V → W ⊆ U ∩ V.
+/3/. nqed.
+
+nlemma cupC : ∀S. ∀a,b:Ω^S.a ∪ b = b ∪ a.
+#S a b; @; #w; *; nnormalize; /2/; nqed.
+
+nlemma cupID : ∀S. ∀a:Ω^S.a ∪ a = a.
+#S a; @; #w; ##[*; //] /2/; nqed.
+
+(* XXX Bug notazione \cup, niente parentesi *)
+nlemma cupA : ∀S.∀a,b,c:Ω^S.a ∪ b ∪ c = a ∪ (b ∪ c).
+#S a b c; @; #w; *; /3/; *; /3/; nqed.
+
+ndefinition Empty_set : ∀A.Ω^A ≝ λA.{ x | False }.
+
+notation "∅" non associative with precedence 90 for @{ 'empty }.
+interpretation "empty set" 'empty = (Empty_set ?).
+
+nlemma cup0 :∀S.∀A:Ω^S.A ∪ ∅ = A.
+#S p; @; #w; ##[*; //| #; @1; //] *; nqed.
+