]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/relations_to_o-algebra.ma
some minor fixes
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / relations_to_o-algebra.ma
index b3939f90ba68dfd1f1fa57eed6e8003d322f21e9..af3015eb368737d5f62f3ce5c8689399edfc9e07 100644 (file)
@@ -15,7 +15,7 @@
 include "relations.ma".
 include "o-algebra.ma".
 
-definition SUBSETS: objs1 SET → OAlgebra.
+definition POW': objs1 SET → OAlgebra.
  intro A; constructor 1;
   [ apply (Ω \sup A);
   | apply subseteq;
@@ -42,16 +42,16 @@ definition SUBSETS: objs1 SET → OAlgebra.
   | intros; split; intro;
      [ cases f; cases x1; exists [apply w1] exists [apply w] assumption;
      | cases e; cases x; exists; [apply w1] [ assumption | exists; [apply w] assumption]]
-  | intros; intros 2; cases (f (singleton ? a) ?);
+  | intros; intros 2; cases (f {(a)} ?); 
      [ exists; [apply a] [assumption | change with (a = a); apply refl1;]
      | change in x1 with (a = w); change with (mem A a q); apply (. (x1‡#));
        assumption]]
 qed.
 
-definition powerset_of_SUBSETS: ∀A.oa_P (SUBSETS A) → Ω \sup A ≝ λA,x.x.
-coercion powerset_of_SUBSETS.
+definition powerset_of_POW': ∀A.oa_P (POW' A) → Ω \sup A ≝ λA,x.x.
+coercion powerset_of_POW'.
 
-definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (SUBSETS o1) (SUBSETS o2).
+definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (POW' o1) (POW' o2).
  intros;
  constructor 1;
   [ constructor 1; 
@@ -112,7 +112,7 @@ lemma orelation_of_relation_preserves_equality:
 qed.
 
 lemma orelation_of_relation_preserves_identity:
- ∀o1:REL. eq2 ? (orelation_of_relation ?? (id1 ? o1)) (id2 OA (SUBSETS o1)).
+ ∀o1:REL. eq2 ? (orelation_of_relation ?? (id1 ? o1)) (id2 OA (POW' o1)).
  intros; split; intro; split; whd; intro; 
   [ change with ((∀x. x ♮(id1 REL o1) a1→x∈a) → a1 ∈ a); intros;
     apply (f a1); change with (a1 = a1); apply refl1;
@@ -141,7 +141,7 @@ alias symbol "compose" = "category1 composition".
 lemma orelation_of_relation_preserves_composition:
  ∀o1,o2,o3:REL.∀F: arrows1 ? o1 o2.∀G: arrows1 ? o2 o3.
   orelation_of_relation ?? (G ∘ F) =
-   comp2 OA (SUBSETS o1) (SUBSETS o2) (SUBSETS o3)
+   comp2 OA (POW' o1) (POW' o2) (POW' o3)
     ?? (*(orelation_of_relation ?? F) (orelation_of_relation ?? G)*).
  [ apply (orelation_of_relation ?? F); | apply (orelation_of_relation ?? G); ]
  intros; split; intro; split; whd; intro; whd in ⊢ (% → %); intros;
@@ -161,9 +161,9 @@ lemma orelation_of_relation_preserves_composition:
   | cases f1; clear f1; cases x; clear x; apply (f w); assumption;]
 qed.
 
-definition SUBSETS': carr3 (arrows3 CAT2 (category2_of_category1 REL) OA).
+definition POW: carr3 (arrows3 CAT2 (category2_of_category1 REL) OA).
  constructor 1;
-  [ apply SUBSETS;
+  [ apply POW';
   | intros; constructor 1;
      [ apply (orelation_of_relation S T);
      | intros; apply (orelation_of_relation_preserves_equality S T a a' e); ]
@@ -171,22 +171,24 @@ definition SUBSETS': carr3 (arrows3 CAT2 (category2_of_category1 REL) OA).
   | apply orelation_of_relation_preserves_composition; ]
 qed.
 
-theorem SUBSETS_faithful:
+theorem POW_faithful:
  ∀S,T.∀f,g:arrows2 (category2_of_category1 REL) S T.
-  map_arrows2 ?? SUBSETS' ?? f = map_arrows2 ?? SUBSETS' ?? g → f=g.
- intros; unfold SUBSETS' in e; simplify in e; cases e;
+  map_arrows2 ?? POW ?? f = map_arrows2 ?? POW ?? g → f=g.
+ intros; unfold POW in e; simplify in e; cases e;
  unfold orelation_of_relation in e3; simplify in e3; clear e e1 e2 e4;
- intros 2; lapply (e3 (singleton ? x)); cases Hletin;
+ intros 2; cases (e3 {(x)}); 
  split; intro; [ lapply (s y); | lapply (s1 y); ]
   [2,4: exists; [1,3:apply x] split; [1,3: assumption |*: change with (x=x); apply rule #]
-  |*: cases Hletin1; cases x1; change in f3 with (eq1 ? x w); apply (. f3‡#); assumption;]
+  |*: cases Hletin; cases x1; change in f3 with (x =_1 w); apply (. f3‡#); assumption;]
 qed.
 
+interpretation "lifting singl" 'singl x = 
+ (fun11 _ (objs2 (POW _))  (singleton _) x).
 
-theorem SUBSETS_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? SUBSETS' S T g = f).
+theorem POW_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? POW S T g = f).
  intros; exists;
   [ constructor 1; constructor 1;
-     [ apply (λx:carr S.λy:carr T. y ∈ f (singleton S x));
+     [ apply (λx:carr S.λy:carr T. y ∈ f {(x)});
      | intros; unfold FunClass_1_OF_carr2; lapply (.= e1‡#);
         [4: apply mem; |6: apply Hletin;|1,2,3,5: skip]
        lapply (#‡prop11 ?? f ?? (†e)); [6: apply Hletin; |*:skip ]]