]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs_to_o-basic_topologies.ma
few more files, one diverges
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-basic_pairs_to_o-basic_topologies.ma
index b2dfffd02952995536a22a7710c54ff9984c8552..b745652569c9458b389903f2a220c93c03c3d6f9 100644 (file)
@@ -19,12 +19,12 @@ include "o-basic_topologies.ma".
 alias symbol "eq" = "setoid1 eq".
 
 (* Qui, per fare le cose per bene, ci serve la nozione di funtore categorico *)
-definition o_basic_topology_of_o_basic_pair: OBP → BTop.
+definition o_basic_topology_of_o_basic_pair: OBP → OBTop.
  intro t;
  constructor 1;
   [ apply (Oform t);
-  | apply (□_t ∘ Ext⎽t);
-  | apply (◊_t ∘ Rest⎽t);
+  | apply (□t ∘ Ext⎽t);
+  | apply (◊t ∘ Rest⎽t);
   | apply hide; intros 2; split; intro;
      [ change with ((⊩) \sup ⎻* ((⊩) \sup ⎻ U) ≤ (⊩) \sup ⎻* ((⊩) \sup ⎻ V));
        apply (. (#‡(lemma_10_4_a ?? (⊩) V)^-1));
@@ -38,7 +38,7 @@ definition o_basic_topology_of_o_basic_pair: OBP → BTop.
           apply (. (or_prop2 : ?) ^ -1);
           apply oa_leq_refl; ]]
   | apply hide; intros 2; split; intro;
-     [ change with (◊_t ((⊩) \sup * U) ≤ ◊_t ((⊩) \sup * V));
+     [ change with (◊⎽t ((⊩) \sup * U) ≤ ◊⎽t ((⊩) \sup * V));
        apply (. ((lemma_10_4_b ?? (⊩) U)^-1)‡#);
        apply (f_image_monotone ?? (⊩) ? ((⊩)* V));
        apply f_star_image_monotone;
@@ -51,7 +51,7 @@ definition o_basic_topology_of_o_basic_pair: OBP → BTop.
           apply oa_leq_refl; ]]
   | apply hide; intros;
     apply (.= (oa_overlap_sym' : ?));
-    change with ((◊_t ((⊩)* V) >< (⊩)⎻* ((⊩)⎻ U)) = (U >< (◊_t ((⊩)* V))));
+    change with ((◊⎽t ((⊩)* V) >< (⊩)⎻* ((⊩)⎻ U)) = (U >< (◊⎽t ((⊩)* V))));
     apply (.= (or_prop3 ?? (⊩) ((⊩)* V) ?));
     apply (.= #‡(lemma_10_3_a : ?));
     apply (.= (or_prop3 : ?)^-1);
@@ -60,7 +60,7 @@ qed.
 
 definition o_continuous_relation_of_o_relation_pair:
  ∀BP1,BP2.arrows2 OBP BP1 BP2 →
-  arrows2 BTop (o_basic_topology_of_o_basic_pair BP1) (o_basic_topology_of_o_basic_pair BP2).
+  arrows2 OBTop (o_basic_topology_of_o_basic_pair BP1) (o_basic_topology_of_o_basic_pair BP2).
  intros (BP1 BP2 t);
  constructor 1;
   [ apply (t \sub \f);
@@ -93,15 +93,15 @@ definition o_continuous_relation_of_o_relation_pair:
 qed.
 
 
-definition OR : carr3 (arrows3 CAT2 OBP BTop).
+definition OR : carr3 (arrows3 CAT2 OBP OBTop).
 constructor 1;
 [ apply o_basic_topology_of_o_basic_pair;
 | intros; constructor 1;
   [ apply o_continuous_relation_of_o_relation_pair;
   | apply hide; 
     intros; whd; unfold o_continuous_relation_of_o_relation_pair; simplify;;
-    change with ((a \sub \f ⎻* ∘ A (o_basic_topology_of_o_basic_pair S)) =
-                 (a' \sub \f ⎻*∘A (o_basic_topology_of_o_basic_pair S)));
+    change with ((a \sub \f ⎻* ∘ oA (o_basic_topology_of_o_basic_pair S)) =
+                 (a' \sub \f ⎻*∘ oA (o_basic_topology_of_o_basic_pair S)));
     whd in e; cases e; clear e e2 e3 e4;
     change in ⊢ (? ? ? (? ? ? ? ? % ?) ?) with ((⊩\sub S)⎻* ∘ (⊩\sub S)⎻);
     apply (.= (comp_assoc2 ? ???? ?? a\sub\f⎻* ));
@@ -113,7 +113,7 @@ constructor 1;
     change in ⊢ (? ? ? (? ? ? ? ? ? %) ?) with (a'\sub\f⎻* ∘ (⊩\sub S)⎻* );    
     apply (.= (comp_assoc2 ? ???? ?? a'\sub\f⎻* )^-1);
     apply refl2;]
-| intros 2 (o a); apply rule #;
+| intros 2 (o a); apply refl1;
 | intros 6; apply refl1;]
 qed.