X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fformal_topology%2Foverlap%2Fo-concrete_spaces.ma;h=90086b4d2f0e8646b973ef60925f8ccaee6db19b;hb=21cb6b92e615f08280e9f032718593b4a6a1d3a5;hp=57ec577aec9d1248c0f70e521dc62329a4dafbeb;hpb=05669a3646cbf8219bf772718b4b2c5bfd1ddd82;p=helm.git diff --git a/helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma b/helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma index 57ec577ae..90086b4d2 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma @@ -15,134 +15,120 @@ include "o-basic_pairs.ma". include "o-saturations.ma". -definition A : ∀b:BP. unary_morphism1 (form b) (form b). +definition A : ∀b:OBP. unary_morphism1 (Oform b) (Oform b). intros; constructor 1; [ apply (λx.□_b (Ext⎽b x)); - | do 2 unfold FunClass_1_OF_Type_OF_setoid21; intros; apply (†(†e));] + | 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 : Type2 ≝ - { bp:> BP; +record Oconcrete_space : Type2 ≝ + { Obp:> OBP; (*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 + Odownarrow: unary_morphism1 (Oform Obp) (Oform Obp); + Odownarrow_is_sat: is_o_saturation ? Odownarrow; + Oconverges: ∀q1,q2. + (Ext⎽Obp q1 ∧ (Ext⎽Obp q2)) = (Ext⎽Obp ((Odownarrow q1) ∧ (Odownarrow q2))); + Oall_covered: Ext⎽Obp (oa_one (Oform Obp)) = oa_one (Oconcr Obp); + Oil2: ∀I:SET.∀p:arrows2 SET1 I (Oform Obp). + Odownarrow (∨ { x ∈ I | Odownarrow (p x) | down_p ???? }) = + ∨ { x ∈ I | Odownarrow (p x) | down_p ???? }; + Oil1: ∀q.Odownarrow (A ? q) = A ? q }. interpretation "o-concrete space downarrow" 'downarrow x = - (fun11 __ (downarrow _) x). + (fun11 __ (Odownarrow _) x). -definition bp': concrete_space → basic_pair ≝ λc.bp c. -coercion bp'. - -definition bp'': concrete_space → objs2 BP. - intro; apply (bp' c); -qed. -coercion bp''. - -definition binary_downarrow : - ∀C:concrete_space.binary_morphism1 (form C) (form C) (form C). +definition Obinary_downarrow : + ∀C:Oconcrete_space.binary_morphism1 (Oform C) (Oform C) (Oform C). intros; constructor 1; -[ intros; apply (↓ t ∧ ↓ t1); +[ 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). +interpretation "concrete_space binary ↓" 'fintersects a b = (fun21 _ _ _ (Obinary_downarrow _) a b). -record convergent_relation_pair (CS1,CS2: concrete_space) : Type2 ≝ - { rp:> arrows2 ? CS1 CS2; - respects_converges: - ∀b,c. eq1 ? (rp\sub\c⎻ (Ext⎽CS2 (b ↓ c))) (Ext⎽CS1 (rp\sub\f⎻ b ↓ rp\sub\f⎻ c)); - respects_all_covered: - eq1 ? (rp\sub\c⎻ (Ext⎽CS2 (oa_one (form CS2)))) - (Ext⎽CS1 (oa_one (form CS1))) +record Oconvergent_relation_pair (CS1,CS2: Oconcrete_space) : Type2 ≝ + { Orp:> arrows2 ? CS1 CS2; + Orespects_converges: + ∀b,c. eq1 ? (Orp\sub\c⎻ (Ext⎽CS2 (b ↓ c))) (Ext⎽CS1 (Orp\sub\f⎻ b ↓ Orp\sub\f⎻ c)); + Orespects_all_covered: + eq1 ? (Orp\sub\c⎻ (Ext⎽CS2 (oa_one (Oform CS2)))) + (Ext⎽CS1 (oa_one (Oform 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 → setoid2. - intros; +definition Oconvergent_relation_space_setoid: Oconcrete_space → Oconcrete_space → setoid2. + intros (c c1); constructor 1; - [ apply (convergent_relation_pair c c1) + [ apply (Oconvergent_relation_pair c c1) | constructor 1; - [ intros; - apply (relation_pair_equality c c1 c2 c3); + [ intros (c2 c3); + apply (Orelation_pair_equality c c1 c2 c3); | intros 1; apply refl2; | intros 2; apply sym2; | intros 3; apply trans2]] qed. +definition Oconvergent_relation_space_of_Oconvergent_relation_space_setoid: + ∀CS1,CS2.carr2 (Oconvergent_relation_space_setoid CS1 CS2) → + Oconvergent_relation_pair CS1 CS2 ≝ λP,Q,c.c. +coercion Oconvergent_relation_space_of_Oconvergent_relation_space_setoid. -definition rp'': ∀CS1,CS2.carr2 (convergent_relation_space_setoid CS1 CS2) → arrows2 BP CS1 CS2 ≝ - λCS1,CS2,c.rp ?? c. -coercion rp''. - -definition rp''': ∀CS1,CS2.Type_OF_setoid2 (convergent_relation_space_setoid CS1 CS2) → arrows2 BP CS1 CS2 ≝ - λCS1,CS2,c.rp ?? c. -coercion rp'''. - -definition rp'''': ∀CS1,CS2.Type_OF_setoid2 (convergent_relation_space_setoid CS1 CS2) → carr2 (arrows2 BP CS1 CS2) ≝ - λCS1,CS2,c.rp ?? c. -coercion rp''''. - -definition convergent_relation_space_composition: - ∀o1,o2,o3: concrete_space. +definition Oconvergent_relation_space_composition: + ∀o1,o2,o3: Oconcrete_space. binary_morphism2 - (convergent_relation_space_setoid o1 o2) - (convergent_relation_space_setoid o2 o3) - (convergent_relation_space_setoid o1 o3). + (Oconvergent_relation_space_setoid o1 o2) + (Oconvergent_relation_space_setoid o2 o3) + (Oconvergent_relation_space_setoid o1 o3). intros; constructor 1; [ intros; whd in t t1 ⊢ %; constructor 1; - [ apply (t1 ∘ t); + [ apply (c1 ∘ c); | intros; - change in ⊢ (? ? ? % ?) with (t\sub\c⎻ (t1\sub\c⎻ (Ext⎽o3 (b↓c)))); - unfold FunClass_1_OF_Type_OF_setoid21; + change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (b↓c2)))); alias symbol "trans" = "trans1". - apply (.= († (respects_converges : ?))); - apply (respects_converges ?? t (t1\sub\f⎻ b) (t1\sub\f⎻ c)); - | change in ⊢ (? ? ? % ?) with (t\sub\c⎻ (t1\sub\c⎻ (Ext⎽o3 (oa_one (form o3))))); - unfold FunClass_1_OF_Type_OF_setoid21; - apply (.= (†(respects_all_covered :?))); - apply rule (respects_all_covered ?? t);] + apply (.= († (Orespects_converges : ?))); + apply (Orespects_converges ?? c (c1\sub\f⎻ b) (c1\sub\f⎻ c2)); + | change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (oa_one (Oform o3))))); + apply (.= (†(Orespects_all_covered :?))); + apply rule (Orespects_all_covered ?? c);] | intros; change with (b ∘ a = b' ∘ a'); - change in e with (rp'' ?? a = rp'' ?? a'); - change in e1 with (rp'' ?? b = rp ?? b'); + change in e with (Orp ?? a = Orp ?? a'); + change in e1 with (Orp ?? b = Orp ?? b'); apply (e‡e1);] qed. -definition CSPA: category2. +definition OCSPA: category2. constructor 1; - [ apply concrete_space - | apply convergent_relation_space_setoid + [ apply Oconcrete_space + | apply Oconvergent_relation_space_setoid | intro; constructor 1; [ apply id2 | intros; apply refl1; | apply refl1] - | apply convergent_relation_space_composition + | apply Oconvergent_relation_space_composition | intros; simplify; whd in a12 a23 a34; change with (a34 ∘ (a23 ∘ a12) = (a34 ∘ a23) ∘ a12); apply rule ASSOC; | intros; simplify; - change with (a ∘ id2 ? o1 = a); + change with (a ∘ id2 OBP o1 = a); apply (id_neutral_right2 : ?); | intros; simplify; change with (id2 ? o2 ∘ a = a); apply (id_neutral_left2 : ?);] qed. + +definition Oconcrete_space_of_OCSPA : objs2 OCSPA → Oconcrete_space ≝ λx.x. +coercion Oconcrete_space_of_OCSPA. + +definition Oconvergent_relation_space_setoid_of_arrows2_OCSPA : + ∀P,Q. arrows2 OCSPA P Q → Oconvergent_relation_space_setoid P Q ≝ λP,Q,x.x. +coercion Oconvergent_relation_space_setoid_of_arrows2_OCSPA. +