]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/nlibrary/sets/sets.ma
Some refactoring in set*.ma, some new notations and new hints for \cup.
[helm.git] / helm / software / matita / nlibrary / sets / sets.ma
index 5600b9a16a49ed4fc7992de19a5fed5de94dba0f..544db6a562bb4cd266d38082ace2ac0dc6197fc0 100644 (file)
@@ -49,11 +49,8 @@ nlemma subseteq_trans: ∀A.∀S,T,U: Ω^A. S ⊆ T → T ⊆ U → S ⊆ U.
 include "properties/relations1.ma".
 
 ndefinition seteq: ∀A. equivalence_relation1 (Ω^A).
- #A; @
-  [ napply (λS,S'. S ⊆ S' ∧ S' ⊆ S)
-  | /2/
-  | #S; #S'; *; /3/
-  | #S; #T; #U; *; #H1; #H2; *; /4/]
+#A; @(λS,S'. S ⊆ S' ∧ S' ⊆ S); /2/; ##[ #A B; *; /3/]
+#S T U; *; #H1 H2; *; /4/;
 nqed.
 
 include "sets/setoids1.ma".
@@ -68,18 +65,21 @@ nqed.
 include "hints_declaration.ma". 
 
 alias symbol "hint_decl" = "hint_decl_Type2".
-unification hint 0 ≔ A ⊢ carr1 (mk_setoid1 (Ω^A) (eq1 (powerclass_setoid A))) ≡ Ω^A.
+unification hint 0 ≔ A;
+  R ≟ (mk_setoid1 (Ω^A) (eq1 (powerclass_setoid A)))
+(*--------------------------------------------------*)⊢ 
+     carr1 R ≡ Ω^A.
 
 (************ SETS OVER SETOIDS ********************)
 
 include "logic/cprop.ma".
 
-nrecord ext_powerclass (A: setoid) : Type[1] ≝
{ ext_carr:> Ω^A; (* qui pc viene dichiarato con un target preciso... 
-                forse lo si vorrebbe dichiarato con un target più lasco 
-                ma la sintassi :> non lo supporta *)
+nrecord ext_powerclass (A: setoid) : Type[1] ≝ { 
  ext_carr:> Ω^A; (* qui pc viene dichiarato con un target preciso... 
+                      forse lo si vorrebbe dichiarato con un target più lasco 
+                      ma la sintassi :> non lo supporta *)
    ext_prop: ∀x,x':A. x=x' → (x ∈ ext_carr) = (x' ∈ ext_carr) 
- }.
+}.
  
 notation > "𝛀 ^ term 90 A" non associative with precedence 70 
 for @{ 'ext_powerclass $A }.
@@ -107,17 +107,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 (Ω^?). 
-*)
-
 nlemma mem_ext_powerclass_setoid_is_morph: 
- ∀A. unary_morphism1 (setoid1_of_setoid A) (unary_morphism1_setoid1 (ext_powerclass_setoid A) CPROP).
+ ∀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/.
@@ -126,65 +117,72 @@ nqed.
 unification hint 0 ≔  AA, x, S;  
      A ≟ carr AA,
      SS ≟ (ext_carr ? S),
-     TT ≟ (mk_unary_morphism1  
+     TT ≟ (mk_unary_morphism1 ?? 
              (λx:setoid1_of_setoid ?.
-               mk_unary_morphism1 
+               mk_unary_morphism1 ??
                  (λS:ext_powerclass_setoid ?. x ∈ S)
-                 (prop11  (mem_ext_powerclass_setoid_is_morph AA x)))
-             (prop11  (mem_ext_powerclass_setoid_is_morph AA))),
+                 (prop11 ?? (mem_ext_powerclass_setoid_is_morph AA x)))
+             (prop11 ?? (mem_ext_powerclass_setoid_is_morph AA))),
      XX ≟ (ext_powerclass_setoid AA)
   (*-------------------------------------*) ⊢ 
       fun11 (setoid1_of_setoid AA)
        (unary_morphism1_setoid1 XX CPROP) TT x S 
     ≡ mem A SS x.
 
-nlemma subseteq_is_morph: ∀A. unary_morphism1 (ext_powerclass_setoid A)
- (unary_morphism1_setoid1 (ext_powerclass_setoid A) CPROP).
+nlemma set_ext : ∀S.∀A,B:Ω^S.A =_1 B → ∀x:S.(x ∈ A) =_1 (x ∈ B).
+#S A B; *; #H1 H2 x; @; ##[ napply H1 | napply H2] nqed.
+
+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).
  #A; napply (mk_binary_morphism1 … (λS,S':𝛀^A. S ⊆ S'));
  #a; #a'; #b; #b'; *; #H1; #H2; *; /5/ by mk_iff, sym1, subseteq_trans;
 nqed.
 
 alias symbol "hint_decl" (instance 1) = "hint_decl_Type2".
-unification hint 0 ≔ A,a,a'
- (*-----------------------------------------------------------------*) ⊢
-  eq_rel ? (eq0 A) a a' ≡ eq_rel1 ? (eq1 (setoid1_of_setoid A)) a a'.
+unification hint 0 ≔ A,x,y
+(*-----------------------------------------------*) ⊢
+   eq_rel ? (eq0 A) x y ≡ eq_rel1 ? (eq1 (setoid1_of_setoid A)) x y.
+   
+(* XXX capire come mai questa hint non funziona se porto su (setoid1_of_setoid A) *)
 
 nlemma intersect_is_ext: ∀A. 𝛀^A → 𝛀^A → 𝛀^A.
- #A; #S; #S'; @ (S ∩ S');
- #a; #a'; #Ha; @; *; #H1; #H2; @ 
-  [##1,2: napply (. Ha^-1‡#); nassumption;
-##|##3,4: napply (. Ha‡#); nassumption]
+#S A B; @ (A ∩ B); #x y Exy; @; *; #H1 H2; @;
+##[##1,2: napply (. Exy^-1‡#); nassumption;
+##|##3,4: napply (. Exy‡#); nassumption]
 nqed.
 
 alias symbol "hint_decl" = "hint_decl_Type1".
 unification hint 0 ≔ 
   A : setoid, B,C : ext_powerclass A;
-  R ≟ (mk_ext_powerclass ? (B ∩ C) (ext_prop ? (intersect_is_ext ? B C)))
-  
+  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. unary_morphism1 (powerclass_setoid A) (unary_morphism1_setoid1 (powerclass_setoid A) (powerclass_setoid A)).
+ ∀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/.
 nqed.
 
 alias symbol "hint_decl" = "hint_decl_Type1".
-unification hint 0 ≔ 
-  A : Type[0], B,C : Ω^A;
-  R ≟ (mk_unary_morphism1 …
-       (λS. mk_unary_morphism1 … (λS'.S ∩ S') (prop11 … (intersect_is_morph A S))) 
-       (prop11 … (intersect_is_morph A)))
-   ⊢ 
-    R B C  ≡ intersect ? B C.
+unification hint 0 ≔ A : Type[0], B,C : Ω^A;
+  R ≟ mk_unary_morphism1 ??
+       (λS. mk_unary_morphism1 ?? (λS'.S ∩ S') (prop11 ?? (intersect_is_morph A S))) 
+       (prop11 ?? (intersect_is_morph A))
+(*------------------------------------------------------------------------*) ⊢ 
+    fun11 ?? (fun11 ?? R B) C  ≡ intersect A B C.
 
 interpretation "prop21 ext" 'prop2 l r =
  (prop11 (ext_powerclass_setoid ?)
   (unary_morphism1_setoid1 (ext_powerclass_setoid ?) ?) ? ?? l ?? r).
 
 nlemma intersect_is_ext_morph: 
- ∀A. unary_morphism1 (ext_powerclass_setoid A)
+ ∀A. 
+ (ext_powerclass_setoid A) ⇒_1
   (unary_morphism1_setoid1 (ext_powerclass_setoid A) (ext_powerclass_setoid A)).
  #A; napply (mk_binary_morphism1 … (intersect_is_ext …));
  #a; #a'; #b; #b'; #Ha; #Hb; napply (prop11 … (intersect_is_morph A)); nassumption.
@@ -193,18 +191,73 @@ nqed.
 unification hint 1 ≔ 
       AA : setoid, B,C : 𝛀^AA;
       A ≟ carr AA,
-      R ≟ (mk_unary_morphism1 
-              (λS:ext_powerclass_setoid AA.
+      R ≟ (mk_unary_morphism1 ??
+              (λS:𝛀^AA.
                mk_unary_morphism1 ??
-                (λS':ext_powerclass_setoid AA.
+                (λS':𝛀^AA.
                   mk_ext_powerclass AA (S∩S') (ext_prop AA (intersect_is_ext ? S S')))
-                (prop11  (intersect_is_ext_morph AA S))) 
-              (prop11  (intersect_is_ext_morph AA))) ,
+                (prop11 ?? (intersect_is_ext_morph AA S))) 
+              (prop11 ?? (intersect_is_ext_morph AA))) ,
        BB ≟ (ext_carr ? B),
        CC ≟ (ext_carr ? C)
    (* ------------------------------------------------------*) ⊢ 
             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*)
+#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 (?∨?);
+napply (.= (set_ext ??? EA x)‡#);
+napply (.= #‡(set_ext ??? EB x)); //;
+nqed.
+
+nlemma union_is_ext: ∀A. 𝛀^A → 𝛀^A → 𝛀^A.
+ #S A B; @ (A ∪ B); #x y Exy; @; *; #H1; 
+##[##1,3: @; ##|##*: @2 ]
+##[##1,3: napply (. (Exy^-1)╪_1#) 
+##|##2,4: napply (. Exy╪_1#)]
+nassumption;
+nqed.
+
+alias symbol "hint_decl" = "hint_decl_Type1".
+unification hint 0 ≔
+   A : setoid, B,C :  𝛀^A;
+   R ≟ (mk_ext_powerclass ? (B ∪ C) (ext_prop ? (union_is_ext ? B C)))
+(*-------------------------------------------------------------------------*)  ⊢
+    ext_carr A R ≡ union ? (ext_carr ? B) (ext_carr ? C).
+
+unification hint 0 ≔ S:Type[0], A,B:Ω^S;
+  MM ≟ mk_unary_morphism1 ??
+        (λA.mk_unary_morphism1 ?? (λB.A ∪ B) (prop11 ?? (union_is_morph S A)))
+        (prop11 ?? (union_is_morph 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 *)
+#A; napply (mk_binary_morphism1 …  (union_is_ext …));
+#x1 x2 y1 y2 Ex Ey; napply (prop11 … (union_is_morph A)); nassumption.
+nqed.
+            
+unification hint 1 ≔
+  AA : setoid, B,C : 𝛀^AA;
+  A ≟ carr AA,
+  R ≟ (mk_unary_morphism1 ??
+          (λS:𝛀^AA.
+           mk_unary_morphism1 ??
+            (λS':𝛀^AA.
+              mk_ext_powerclass AA (S ∪ S') (ext_prop AA (union_is_ext ? S S')))
+            (prop11 ?? (union_is_ext_morph AA S)))
+          (prop11 ?? (union_is_ext_morph AA))) ,
+   BB ≟ (ext_carr ? B),
+   CC ≟ (ext_carr ? C)
+(*------------------------------------------------------*) ⊢
+   ext_carr AA (R B C) ≡ union A BB CC.
+
 (*
 alias symbol "hint_decl" = "hint_decl_Type2".
 unification hint 0 ≔
@@ -284,21 +337,20 @@ nqed.
 (******************* first omomorphism theorem for sets **********************)
 
 ndefinition eqrel_of_morphism:
- ∀A,B. unary_morphism A B → compatible_equivalence_relation A.
+ ∀A,B. A ⇒_0 B → compatible_equivalence_relation A.
  #A; #B; #f; @
   [ @ [ napply (λx,y. f x = f y) ] /2/;
 ##| #x; #x'; #H; nwhd; alias symbol "prop1" = "prop1".
 napply (.= (†H)); // ]
 nqed.
 
-ndefinition canonical_proj: ∀A,R. unary_morphism A (quotient A R).
+ndefinition canonical_proj: ∀A,R. A ⇒_0 (quotient A R).
  #A; #R; @
   [ napply (λx.x) |  #a; #a'; #H; napply (compatibility … R … H) ]
 nqed.
 
 ndefinition quotiented_mor:
- ∀A,B.∀f:unary_morphism A B.
-  unary_morphism (quotient … (eqrel_of_morphism … f)) B.
+ ∀A,B.∀f:A ⇒_0 B.(quotient … (eqrel_of_morphism … f)) ⇒_0 B.
  #A; #B; #f; @ [ napply f ] //.
 nqed.
 
@@ -309,26 +361,26 @@ nlemma first_omomorphism_theorem_functions1:
 
 alias symbol "eq" = "setoid eq".
 ndefinition surjective ≝
- λA,B.λS: ext_powerclass A.λT: ext_powerclass B.λf:unary_morphism A B.
+ λA,B.λS: ext_powerclass A.λT: ext_powerclass B.λf:A ⇒_0 B.
   ∀y. y ∈ T → ∃x. x ∈ S ∧ f x = y.
 
 ndefinition injective ≝
- λA,B.λS: ext_powerclass A.λf:unary_morphism A B.
+ λA,B.λS: ext_powerclass A.λf:A ⇒_0 B.
   ∀x,x'. x ∈ S → x' ∈ S → f x = f x' → x = x'.
 
 nlemma first_omomorphism_theorem_functions2:
- ∀A,B.∀f: unary_morphism A B. 
+ ∀A,B.∀f:A ⇒_0 B. 
    surjective … (Full_set ?) (Full_set ?) (canonical_proj ? (eqrel_of_morphism … f)).
 /3/. nqed.
 
 nlemma first_omomorphism_theorem_functions3:
- ∀A,B.∀f: unary_morphism A B. 
+ ∀A,B.∀f:A ⇒_0 B. 
    injective … (Full_set ?) (quotiented_mor … f).
  #A; #B; #f; nwhd; #x; #x'; #Hx; #Hx'; #K; nassumption.
 nqed.
 
 nrecord isomorphism (A, B : setoid) (S: ext_powerclass A) (T: ext_powerclass B) : Type[0] ≝
- { iso_f:> unary_morphism A B;
+ { iso_f:> A ⇒_0 B;
    f_closed: ∀x. x ∈ S → iso_f x ∈ T;
    f_sur: surjective … S T iso_f;
    f_inj: injective … S iso_f