]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs_to_o-basic_topologies.ma
...
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-basic_pairs_to_o-basic_topologies.ma
index bfd76f99e31ed73bd6869051cceb62fd124d0741..1bf31881c29d11fc3b0b1ddb90d96908efd3e552 100644 (file)
@@ -19,7 +19,7 @@ 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);
@@ -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,118 +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.
 
-(*
-axiom DDD : False.
-
-definition sigma_equivalence_relation2:
- ∀C2:CAT2.∀Q.∀X,Y:exT22 ? (λy:C2.Q y).∀P. 
-   equivalence_relation2 (exT22 ? (λf:arrows2 C2 (\fst X) (\fst Y).P f)).
-intros; constructor 1;
-    [ intros(F G); apply (\fst F =_2 \fst G);
-    | intro; apply refl2;
-    | intros 3; apply sym2; assumption;
-    | intros 5; apply (trans2 ?? ??? x1 x2);]
-qed.     
-
-definition Apply : ∀C1,C2: CAT2.arrows3 CAT2 C1 C2 → CAT2.
-intros (C1 C2 F);
-constructor 1; 
-[ apply (exT22 ? (λx:C2.exT22 ? (λy:C1.map_objs2 ?? F y =_\ID x)));
-| intros (X Y); constructor 1;
-  [ apply (exT22 ? (λf:arrows2 C2 (\fst X) (\fst Y).
-           exT22 ? (λg:arrows2 C1 (\fst (\snd X)) (\fst (\snd Y)). 
-           ? (map_arrows2 ?? F ?? g) = f)));
-    intro; apply hide; clear g f; cases X in c; cases Y; cases x; cases x1; clear X Y x x1;
-    simplify; cases H; cases H1; intros; assumption;
-  | apply sigma_equivalence_relation2;] 
-| intro o; constructor 1; 
-   [ apply (id2 C2 (\fst o))
-   | exists[apply (id2 C1 (\fst (\snd o)))] 
-     cases o; cases x; cases H; unfold hide; simplify;
-     apply (respects_id2 ?? F);]
-| intros (o1 o2 o3); constructor 1;
-  [ intros (f g); whd in f g; constructor 1;
-    [ apply (comp2 C2 (\fst o1) (\fst o2) (\fst o3) (\fst f) (\fst g));
-    | exists[apply (comp2 C1 (\fst (\snd o1)) (\fst (\snd o2)) (\fst (\snd o3)) (\fst (\snd f)) (\fst (\snd g)))]
-    cases o1; cases x; cases H; 
-
-(* scrivo gli statement qua cosi' verra' un conflitto :-)
-
-1. definire il funtore OR
-2. dimostrare che ORel e' faithful
-
-3. Definire la funzione
-    Apply:
-     \forall C1,C2: CAT2.  F: arrows3 CAT2 C1 C2 -> CAT2
-    :=
-     constructor 1;
-      [ gli oggetti sono gli oggetti di C1 mappati da F
-      | i morfismi i morfismi di C1 mappati da F
-      | ....
-      ]
-   
-   E : objs CATS === Σx.∃y. F y = x
-  
-   Quindi (Apply C1 C2 F) (che usando da ora in avanti una coercion
-   scrivero' (F C1) ) e' l'immagine di C1 tramite F ed e'
-   una sottocategoria di C2 (qualcosa da dimostare qui??? vedi sotto
-   al punto 5)
-
-4. Definire rOBP (le OBP rappresentabili) come (BP_to_OBP BP)
-  [Si puo' fare lo stesso per le OA: rOA := Rel_to_OA REL ]
-
-5. Dimostrare che OR (il funtore faithful da OBP a OBTop) e' full
-   quando applicato a rOBP.
-   Nota: puo' darsi che faccia storie ad accettare lo statement.
-   Infatti rOBP e' (BP_to_OBP BP) ed e' "una sottocategoria di OBP"
-   e OR va da OBP a OBTop. Non so se tipa subito o se devi dare
-   una "proiezione" da rOBP a OBP.
-
-6. Definire rOBTop come (OBP_to_OBTop rOBP).
-
-7. Per composizione si ha un funtore full and faithful da BP a rOBTop:
-   basta prendere (OR \circ BP_to_OBP).
-
-8. Dimostrare (banale: quasi tutti i campi sono per conversione) che
-   esiste un funtore da rOBTop a BTop. Dimostrare che tale funtore e'
-   faithful e full (banale: tutta conversione).
-
-9. Per composizione si ha un funtore full and faithful da BP a BTop.
-
-10. Dimostrare che i seguenti funtori sono anche isomorphism-dense
-    (http://planetmath.org/encyclopedia/DenseFunctor.html):
-
-    BP_to_OBP
-    OBP_to_OBTop quando applicato alle rOBP
-    OBTop_to_BTop quando applicato alle rOBTop
-
-    Concludere per composizione che anche il funtore da BP a BTop e'
-    isomorphism-dense.
-
-====== Da qui in avanti non e' "necessario" nulla:
-
-== altre cose mancanti
-
-11. Dimostrare che le r* e le * orrizzontali
-    sono isomorfe dando il funtore da r* a * e dimostrando che componendo i
-    due funtori ottengo l'identita'
-
-12. La definizione di r* fa schifo: in pratica dici solo come ottieni
-    qualcosa, ma non come lo caratterizzeresti. Ora un teorema carino
-    e' che una a* (e.g. una aOBP) e' sempre una rOBP dove "a" sta per
-    atomic. Dimostrarlo per tutte le r*.
-
-== categorish/future works
-
-13. definire astrattamente la FG-completion e usare quella per
-    ottenere le BP da Rel e le OBP da OA.
-
-14. indebolire le OA, generalizzare le costruzioni, etc. come detto
-    con Giovanni
-
-*)
-