]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma
SUBSETS_full up to universe inconsistency
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-concrete_spaces.ma
index 135dae3c981e0ae95c48d4c9544b5a059c65451b..d7e0bf649754b0e995b3e791585e729de854a1a4 100644 (file)
 include "o-basic_pairs.ma".
 include "o-saturations.ma".
 
-lemma xxx : ∀x.carr x → carr1 (setoid1_of_setoid x). intros; assumption; qed.
-coercion xxx.
+definition A : ∀b:BP. unary_morphism1 (form b) (form b).
+intros; constructor 1;
+ [ apply (λx.□_b (Ext⎽b x));
+ | intros; apply  (†(†e));]
+qed.
+
+lemma down_p : ∀S:SET1.∀I:SET.∀u:S⇒S.∀c:arrows2 SET1 I S.∀a:I.∀a':I.a=a'→u (c a)=u (c a').
+intros; apply (†(†e));
+qed.
 
-record concrete_space : Type ≝
+record concrete_space : Type2 ≝
  { bp:> BP;
-   downarrow: form bp → oa_P (form bp);
-   downarrow_is_sat: is_saturation ? downarrow;
-   converges: ∀q1,q2:form bp.
-     or_f_minus ?? (⊩) q1 ∧ or_f_minus ?? (⊩) q2 = 
-     or_f_minus ?? (⊩) ((downarrow q1) ∧ (downarrow q2));
-   all_covered: (*⨍^-_bp*) or_f_minus ?? (⊩) (oa_one (form bp)) = oa_one (concr bp);
-   il2: ∀I:setoid.∀p:ums I (oa_P (form bp)).
-     downarrow (oa_join ? I (mk_unary_morphism ?? (λi:I. downarrow (p i)) ?)) =
-     oa_join ? I (mk_unary_morphism ?? (λi:I. downarrow (p i)) ?)
+   (*distr : is_distributive (form bp);*)
+   downarrow: unary_morphism1 (form bp) (form bp);
+   downarrow_is_sat: is_o_saturation ? downarrow;
+   converges: ∀q1,q2.
+     (Ext⎽bp q1 ∧ (Ext⎽bp q2)) = (Ext⎽bp ((downarrow q1) ∧ (downarrow q2)));
+   all_covered: Ext⎽bp (oa_one (form bp)) = oa_one (concr bp);
+   il2: ∀I:SET.∀p:arrows2 SET1 I (form bp).
+     downarrow (∨ { x ∈ I | downarrow (p x) | down_p ???? }) =
+     ∨ { x ∈ I | downarrow (p x) | down_p ???? };
+   il1: ∀q.downarrow (A ? q) = A ? q
  }.
 
-definition bp': concrete_space → basic_pair ≝ λc.bp c.
+interpretation "o-concrete space downarrow" 'downarrow x = 
+  (fun11 __ (downarrow _) x).
 
-coercion bp'.
+definition binary_downarrow : 
+  ∀C:concrete_space.binary_morphism1 (form C) (form C) (form C).
+intros; constructor 1;
+[ intros; apply (↓ c ∧ ↓ c1);
+| intros;
+  alias symbol "prop2" = "prop21".
+  alias symbol "prop1" = "prop11".
+  apply ((†e)‡(†e1));]
+qed.
+
+interpretation "concrete_space binary ↓" 'fintersects a b = (fun21 _ _ _ (binary_downarrow _) a b).
 
-record convergent_relation_pair (CS1,CS2: concrete_space) : Type ≝
- { rp:> arrows1 ? CS1 CS2;
+record convergent_relation_pair (CS1,CS2: concrete_space) : Type2 ≝
+ { rp:> arrows2 ? CS1 CS2;
    respects_converges:
-    ∀b,c.
-     extS ?? rp \sub\c (BPextS CS2 (b ↓ c)) =
-     BPextS CS1 ((extS ?? rp \sub\f b) ↓ (extS ?? rp \sub\f c));
+    ∀b,c. eq1 ? (rp\sub\c⎻ (Ext⎽CS2 (b ↓ c))) (Ext⎽CS1 (rp\sub\f⎻ b ↓ rp\sub\f⎻ c));
    respects_all_covered:
-    extS ?? rp\sub\c (BPextS CS2 (form CS2)) = BPextS CS1 (form CS1)
+     eq1 ? (rp\sub\c⎻ (Ext⎽CS2 (oa_one (form CS2))))
+           (Ext⎽CS1 (oa_one (form CS1)))
  }.
 
-definition rp' : ∀CS1,CS2. convergent_relation_pair CS1 CS2 → relation_pair CS1 CS2 ≝
- λCS1,CS2,c. rp CS1 CS2 c.
-coercion rp'.
-
-definition convergent_relation_space_setoid: concrete_space → concrete_space → setoid1.
+definition convergent_relation_space_setoid: concrete_space → concrete_space → setoid2.
  intros;
  constructor 1;
   [ apply (convergent_relation_pair c c1)
   | constructor 1;
      [ intros;
        apply (relation_pair_equality c c1 c2 c3);
-     | intros 1; apply refl1;
-     | intros 2; apply sym1
-     | intros 3; apply trans1]]
+     | intros 1; apply refl2;
+     | intros 2; apply sym2
+     | intros 3; apply trans2]]
 qed.
 
-definition rp'': ∀CS1,CS2.convergent_relation_space_setoid CS1 CS2 → arrows1 BP CS1 CS2 ≝
- λCS1,CS2,c.rp ?? c.
-
-coercion rp''.
+definition convergent_relation_space_of_convergent_relation_space_setoid: 
+  ∀CS1,CS2.carr2 (convergent_relation_space_setoid CS1 CS2) → 
+    convergent_relation_pair CS1 CS2  ≝ λP,Q,c.c.
+coercion convergent_relation_space_of_convergent_relation_space_setoid.
 
 definition convergent_relation_space_composition:
  ∀o1,o2,o3: concrete_space.
-  binary_morphism1
+  binary_morphism2
    (convergent_relation_space_setoid o1 o2)
    (convergent_relation_space_setoid o2 o3)
    (convergent_relation_space_setoid o1 o3).
  intros; constructor 1;
-     [ intros; whd in c c1 ⊢ %;
+     [ intros; whd in t t1 ⊢ %;
        constructor 1;
-        [ apply (fun1 ??? (comp1 BP ???)); [apply (bp o2) |*: apply rp; assumption]
+        [ apply (c1 ∘ c);
         | intros;
-          change in ⊢ (? ? ? (? ? ? (? ? ? %) ?) ?) with (c1 \sub \c ∘ c \sub \c);
-          change in ⊢ (? ? ? ? (? ? ? ? (? ? ? ? ? (? ? ? (? ? ? %) ?) ?)))
-            with (c1 \sub \f ∘ c \sub \f);
-          change in ⊢ (? ? ? ? (? ? ? ? (? ? ? ? ? ? (? ? ? (? ? ? %) ?))))
-            with (c1 \sub \f ∘ c \sub \f);
-          apply (.= (extS_com ??????));
-          apply (.= (†(respects_converges ?????)));
-          apply (.= (respects_converges ?????));
-          apply (.= (†(((extS_com ??????) \sup -1)‡(extS_com ??????)\sup -1)));
-          apply refl1;
-        | change in ⊢ (? ? ? (? ? ? (? ? ? %) ?) ?) with (c1 \sub \c ∘ c \sub \c);
-          apply (.= (extS_com ??????));
-          apply (.= (†(respects_all_covered ???)));
-          apply (.= respects_all_covered ???);
-          apply refl1]
+          change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (b↓c2))));
+          alias symbol "trans" = "trans1".
+          apply (.= († (respects_converges : ?)));
+          apply (respects_converges ?? c (c1\sub\f⎻ b) (c1\sub\f⎻ c2));
+        | change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (oa_one (form o3)))));
+          apply (.= (†(respects_all_covered :?)));
+          apply rule (respects_all_covered ?? c);]
      | intros;
-       change with (b ∘ a = b' ∘ a');
-       change in H with (rp'' ?? a = rp'' ?? a');
-       change in H1 with (rp'' ?? b = rp ?? b');
-       apply (.= (H‡H1));
-       apply refl1]
+       change with (b ∘ a = b' ∘ a'); 
+       change in e with (rp ?? a = rp ?? a');
+       change in e1 with (rp ?? b = rp ?? b');
+       apply (e‡e1);]
 qed.
 
-definition CSPA: category1.
+definition CSPA: category2.
  constructor 1;
   [ apply concrete_space
   | apply convergent_relation_space_setoid
   | intro; constructor 1;
-     [ apply id1
-     | intros;
-       unfold id; simplify;
-       apply (.= (equalset_extS_id_X_X ??));
-       apply (.= (†((equalset_extS_id_X_X ??)\sup -1‡
-                    (equalset_extS_id_X_X ??)\sup -1)));
-       apply refl1;
-     | apply (.= (equalset_extS_id_X_X ??));
-       apply refl1]
+     [ apply id2
+     | intros; apply refl1;
+     | apply refl1]
   | apply convergent_relation_space_composition
-  | intros; simplify;
+  | intros; simplify; whd in a12 a23 a34;
     change with (a34 ∘ (a23 ∘ a12) = (a34 ∘ a23) ∘ a12);
-    apply (.= ASSOC1);
-    apply refl1
+    apply rule ASSOC;
   | intros; simplify;
-    change with (a ∘ id1 ? o1 = a);
-    apply (.= id_neutral_right1 ????);
-    apply refl1
+    change with (a ∘ id2 BP o1 = a);
+    apply (id_neutral_right2 : ?);
   | intros; simplify;
-    change with (id1 ? o2 ∘ a = a);
-    apply (.= id_neutral_left1 ????);
-    apply refl1]
+    change with (id2 ? o2 ∘ a = a);
+    apply (id_neutral_left2 : ?);]
 qed.
+
+definition concrete_space_of_CSPA : objs2 CSPA → concrete_space ≝ λx.x.
+coercion concrete_space_of_CSPA.
+
+definition convergent_relation_space_setoid_of_arrows2_CSPA :
+ ∀P,Q. arrows2 CSPA P Q → convergent_relation_space_setoid P Q ≝ λP,Q,x.x.
+coercion convergent_relation_space_setoid_of_arrows2_CSPA.
+