]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/relations_to_o-algebra.ma
some notation for map_arrows2
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / relations_to_o-algebra.ma
index af3015eb368737d5f62f3ce5c8689399edfc9e07..8bf57da98bfccfb58e524cd2e8219ae34d79965d 100644 (file)
@@ -17,7 +17,7 @@ include "o-algebra.ma".
 
 definition POW': objs1 SET → OAlgebra.
  intro A; constructor 1;
-  [ apply (Ω \sup A);
+  [ apply (Ω^A);
   | apply subseteq;
   | apply overlaps;
   | apply big_intersects;
@@ -40,7 +40,7 @@ definition POW': objs1 SET → OAlgebra.
      [ assumption
      | whd; intros; cases i; simplify; assumption]
   | intros; split; intro;
-     [ cases f; cases x1; exists [apply w1] exists [apply w] assumption;
+     [ (** screenshot "screen-pow". *) 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 {(a)} ?); 
      [ exists; [apply a] [assumption | change with (a = a); apply refl1;]
@@ -48,10 +48,10 @@ definition POW': objs1 SET → OAlgebra.
        assumption]]
 qed.
 
-definition powerset_of_POW': ∀A.oa_P (POW' A) → Ω \sup A ≝ λA,x.x.
+definition powerset_of_POW': ∀A.oa_P (POW' A) → Ω^A ≝ λA,x.x.
 coercion powerset_of_POW'.
 
-definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (POW' o1) (POW' o2).
+definition orelation_of_relation: ∀o1,o2:REL. o1 ⇒_\r1 o2 → (POW' o1) ⇒_\o2 (POW' o2).
  intros;
  constructor 1;
   [ constructor 1; 
@@ -90,7 +90,8 @@ definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (
 qed.
 
 lemma orelation_of_relation_preserves_equality:
- ∀o1,o2:REL.∀t,t': arrows1 ? o1 o2. t = t' → eq2 ? (orelation_of_relation ?? t) (orelation_of_relation ?? t').
+ ∀o1,o2:REL.∀t,t': o1 ⇒_\r1 o2. 
+   t = t' → orelation_of_relation ?? t =_2 orelation_of_relation ?? t'.
  intros; split; unfold orelation_of_relation; simplify; intro; split; intro;
  simplify; whd in o1 o2;
   [ change with (a1 ∈ minus_star_image ?? t a → a1 ∈ minus_star_image ?? t' a);
@@ -112,7 +113,7 @@ lemma orelation_of_relation_preserves_equality:
 qed.
 
 lemma orelation_of_relation_preserves_identity:
- ∀o1:REL. eq2 ? (orelation_of_relation ?? (id1 ? o1)) (id2 OA (POW' o1)).
+ ∀o1:REL. orelation_of_relation ?? (id1 ? o1) =_2 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;
@@ -139,11 +140,9 @@ qed.
 alias symbol "eq" = "setoid2 eq".
 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 (POW' o1) (POW' o2) (POW' o3)
-    ?? (*(orelation_of_relation ?? F) (orelation_of_relation ?? G)*).
- [ apply (orelation_of_relation ?? F); | apply (orelation_of_relation ?? G); ]
+ ∀o1,o2,o3:REL.∀F: o1 ⇒_\r1 o2.∀G: o2 ⇒_\r1 o3.
+  orelation_of_relation ?? (G ∘ F) = 
+  comp2 OA ??? (orelation_of_relation ?? F) (orelation_of_relation ?? G).
  intros; split; intro; split; whd; intro; whd in ⊢ (% → %); intros;
   [ whd; intros; apply f; exists; [ apply x] split; assumption; 
   | cases f1; clear f1; cases x1; clear x1; apply (f w); assumption;
@@ -173,7 +172,7 @@ qed.
 
 theorem POW_faithful:
  ∀S,T.∀f,g:arrows2 (category2_of_category1 REL) S T.
-  map_arrows2 ?? POW ?? f = map_arrows2 ?? POW ?? g → f=g.
+   POW⎽⇒ f =_2 POW⎽⇒ g → f =_2 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; cases (e3 {(x)}); 
@@ -182,10 +181,12 @@ theorem POW_faithful:
   |*: 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 POW_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? POW S T g = f).
+lemma currify: ∀A,B,C. (A × B ⇒_1 C) → A → (B ⇒_1 C).
+intros; constructor 1; [ apply (b c); | intros; apply (#‡e);]
+qed.
+
+theorem POW_full: ∀S,T.∀f: (POW S) ⇒_\o2 (POW T) . exT22 ? (λg. POW⎽⇒ g = f).
  intros; exists;
   [ constructor 1; constructor 1;
      [ apply (λx:carr S.λy:carr T. y ∈ f {(x)});
@@ -194,17 +195,17 @@ theorem POW_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? POW S T g = f).
        lapply (#‡prop11 ?? f ?? (†e)); [6: apply Hletin; |*:skip ]]  
      | whd; split; whd; intro; simplify; unfold map_arrows2; simplify; 
         [ split;
-           [ change with (∀a1.(∀x. a1 ∈ f (singleton S x) → x ∈ a) → a1 ∈ f⎻* a);
-           | change with (∀a1.a1 ∈ f⎻* a → (∀x.a1 ∈ f (singleton S x) → x ∈ a)); ]
+           [ change with (∀a1.(∀x. a1 ∈ (f {(x):S}) → x ∈ a) → a1 ∈ f⎻* a);
+           | change with (∀a1.a1 ∈ f⎻* a → (∀x.a1 ∈ f {(x):S} → x ∈ a)); ]
         | split;
-           [ change with (∀a1.(∃y:carr T. y ∈ f (singleton S a1) ∧ y ∈ a) → a1 ∈ f⎻ a);
-           | change with (∀a1.a1 ∈ f⎻ a → (∃y:carr T.y ∈ f (singleton S a1) ∧ y ∈ a)); ]
+           [ change with (∀a1.(∃y:carr T. y ∈ f {(a1):S} ∧ y ∈ a) → a1 ∈ f⎻ a);
+           | change with (∀a1.a1 ∈ f⎻ a → (∃y:carr T.y ∈ f {(a1):S} ∧ y ∈ a)); ]
         | split;
-           [ change with (∀a1.(∃x:carr S. a1 ∈ f (singleton S x) ∧ x ∈ a) → a1 ∈ f a);
-           | change with (∀a1.a1 ∈ f a → (∃x:carr S. a1 ∈ f (singleton S x) ∧ x ∈ a)); ]
+           [ change with (∀a1.(∃x:carr S. a1 ∈ f {(x):S} ∧ x ∈ a) → a1 ∈ f a);
+           | change with (∀a1.a1 ∈. f a → (∃x:carr S. a1 ∈ f {(x):S} ∧ x ∈ a)); ]
         | split;
-           [ change with (∀a1.(∀y. y ∈ f (singleton S a1) → y ∈ a) → a1 ∈ f* a);
-           | change with (∀a1.a1 ∈ f* a → (∀y. y ∈ f (singleton S a1) → y ∈ a)); ]]
+           [ change with (∀a1.(∀y. y ∈ f {(a1):S} → y ∈ a) → a1 ∈ f* a);
+           | change with (∀a1.a1 ∈ f* a → (∀y. y ∈ f {(a1):S} → y ∈ a)); ]]
         [ intros; apply ((. (or_prop2 ?? f (singleton ? a1) a)^-1) ? a1);
            [ intros 2; apply (f1 a2); change in f2 with (a2 ∈ f⎻ (singleton ? a1));
              lapply (. (or_prop3 ?? f (singleton ? a2) (singleton ? a1)));