]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/nlibrary/sets/sets.ma
Extending to the nAx set.
[helm.git] / helm / software / matita / nlibrary / sets / sets.ma
index 3d03887ec2e4a163292041486a977f596631b40c..113654ad330de847908ae68b3fcf65faaf9c3033 100644 (file)
@@ -323,6 +323,7 @@ nlemma first_omomorphism_theorem_functions1:
  #A; #B; #f; #x; napply refl;
 nqed.
 
+alias symbol "eq" = "setoid eq".
 ndefinition surjective ≝
  λA,B.λS: ext_powerclass A.λT: ext_powerclass B.λf:unary_morphism A B.
   ∀y. y ∈ T → ∃x. x ∈ S ∧ f x = y.
@@ -351,6 +352,18 @@ 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; *; #xU; #xV; napply H; nassumption;
+nqed.
+
+nlemma subseteq_union_l: ∀A.∀U,V,W:Ω^A.U ⊆ W → V ⊆ W → U ∪ V ⊆ W.
+#A; #U; #V; #W; #H; #H1; #x; *; #Hx; ##[ napply H; ##| napply H1; ##] nassumption;
+nqed. 
+
+nlemma subseteq_intersection_r: ∀A.∀U,V,W:Ω^A.W ⊆ U → W ⊆ V → W ⊆ U ∩ V.
+#A; #U; #V; #W; #H1; #H2; #x; #Hx; @; ##[ napply H1; ##| napply H2; ##] nassumption;
+nqed. 
+
 (*
 nrecord isomorphism (A, B : setoid) (S: qpowerclass A) (T: qpowerclass B) : CProp[0] ≝
  { iso_f:> unary_morphism A B;