X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fformal_topology%2Fconcrete_spaces.ma;h=3c03b4e667803df8815c98e4ef7806ae86b47715;hb=5755ebe4016316a474ad8ab8d33b1a1a9187cc9e;hp=4accb180381603cc35c0b3eb0fd2992bca7e8aef;hpb=7fad6f9727bb6f054c0198cf10354be4b355baef;p=helm.git diff --git a/helm/software/matita/library/formal_topology/concrete_spaces.ma b/helm/software/matita/library/formal_topology/concrete_spaces.ma index 4accb1803..3c03b4e66 100644 --- a/helm/software/matita/library/formal_topology/concrete_spaces.ma +++ b/helm/software/matita/library/formal_topology/concrete_spaces.ma @@ -14,17 +14,107 @@ include "formal_topology/basic_pairs.ma". -interpretation "REL carrier" 'card c = (carrier c). +record concrete_space : Type ≝ + { bp:> BP; + converges: ∀a: concr bp.∀U,V: form bp. a ⊩ U → a ⊩ V → a ⊩ (U ↓ V); + all_covered: ∀x: concr bp. x ⊩ form bp + }. -definition comprehension: ∀b:REL. (b → CProp) → Ω \sup |b| ≝ - λb:REL.λP.{x | ∃p: x ∈ b. P (mk_ssigma ?? x p)}. +definition bp': concrete_space → basic_pair ≝ λc.bp c. -interpretation "subset comprehension" 'comprehension s p = - (comprehension s p). +coercion bp'. -definition ext: ∀o: basic_pair. (form o) → Ω \sup |(concr o)| ≝ - λo,f.{x ∈ (concr o) | x ♮(rel o) f}. +record convergent_relation_pair (CS1,CS2: concrete_space) : Type ≝ + { rp:> arrows1 ? 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)); + respects_all_covered: + extS ?? rp\sub\c (BPextS CS2 (form CS2)) = BPextS CS1 (form CS1) + }. -definition downarrow ≝ - λo: basic_pair.λa,b: form o. - {c | ext ? c ⊆ ext ? a ∩ ext ? b }. \ No newline at end of file +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. + 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]] +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_composition: + ∀o1,o2,o3: concrete_space. + binary_morphism1 + (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 ⊢ %; + constructor 1; + [ apply (fun1 ??? (comp1 BP ???)); [apply (bp o2) |*: apply rp; assumption] + | 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] + | 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] +qed. + +definition CSPA: category1. + 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 convergent_relation_space_composition + | intros; simplify; + change with (a34 ∘ (a23 ∘ a12) = (a34 ∘ a23) ∘ a12); + apply (.= ASSOC1); + apply refl1 + | intros; simplify; + change with (a ∘ id1 ? o1 = a); + apply (.= id_neutral_right1 ????); + apply refl1 + | intros; simplify; + change with (id1 ? o2 ∘ a = a); + apply (.= id_neutral_left1 ????); + apply refl1] +qed.