X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fformal_topology%2Foverlap%2Fo-basic_topologies.ma;h=873a9df60988a632084b8c6f24e9c2b00b619590;hb=a799c56fa883a1318cb42e185c0d0929b368a961;hp=20923337ef6bdea87b4ecd23c8bc49497f9e0478;hpb=c5bbe2a9b9b914f538ae03526c34f2dea5364b1d;p=helm.git diff --git a/helm/software/matita/contribs/formal_topology/overlap/o-basic_topologies.ma b/helm/software/matita/contribs/formal_topology/overlap/o-basic_topologies.ma index 20923337e..873a9df60 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/o-basic_topologies.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/o-basic_topologies.ma @@ -15,12 +15,12 @@ include "o-algebra.ma". include "o-saturations.ma". -record basic_topology: Type ≝ +record basic_topology: Type2 ≝ { carrbt:> OA; A: carrbt ⇒ carrbt; J: carrbt ⇒ carrbt; - A_is_saturation: is_saturation ? A; - J_is_reduction: is_reduction ? J; + A_is_saturation: is_o_saturation ? A; + J_is_reduction: is_o_reduction ? J; compatibility: ∀U,V. (A U >< J V) = (U >< J V) }. @@ -29,7 +29,7 @@ lemma hint: OA → objs2 OA. qed. coercion hint. -record continuous_relation (S,T: basic_topology) : Type ≝ +record continuous_relation (S,T: basic_topology) : Type2 ≝ { cont_rel:> arrows2 OA S T; (* reduces uses eq1, saturated uses eq!!! *) reduced: ∀U. U = J ? U → cont_rel U = J ? (cont_rel U); @@ -108,7 +108,6 @@ theorem continuous_relation_eq_inv': qed. *) -axiom daemon: False. definition continuous_relation_comp: ∀o1,o2,o3. continuous_relation_setoid o1 o2 → @@ -135,11 +134,6 @@ definition continuous_relation_comp: | apply refl1]] qed. -lemma hintx: ∀S,T. (S ⇒ T) → unary_morphism1 S T. - intros; apply t; -qed. -coercion hintx. - definition BTop: category2. constructor 1; [ apply basic_topology @@ -157,22 +151,22 @@ definition BTop: category2. change in e1 with (b⎻* ∘ A o2 = b'⎻* ∘ A o2); apply (.= e‡#); intro x; - change with (eq1 o3 (b⎻* (a'⎻* (A o1 x))) (b'⎻*(a'⎻* (A o1 x)))); - lapply (saturated o1 o2 a' (A o1 x):?) as X; - [ apply ((saturation_idempotent ?? (A_is_saturation o1) x)^-1) ] - change in X with (eq1 ? (a'⎻* (A o1 x)) (A o2 (a'⎻* (A o1 x)))); + change with (b⎻* (a'⎻* (A o1 x)) = b'⎻*(a'⎻* (A o1 x))); alias symbol "trans" = "trans1". alias symbol "prop1" = "prop11". - apply (.= †X); - whd in e1; - lapply (e1 (a'⎻* (A o1 x))) as X1; - change in X1 with (eq1 (oa_P (carrbt o3)) (b⎻* (A o2 (a'⎻* (A o1 x)))) (b'⎻* (A o2 (a' \sup ⎻* (A o1 x))))); - apply (.= X1); alias symbol "invert" = "setoid1 symmetry". - apply (†(X\sup -1));] + lapply (.= †(saturated o1 o2 a' (A o1 x) : ?)); + [3: apply (b⎻* ); | 5: apply Hletin; |1,2: skip; + |apply ((o_saturation_idempotent ?? (A_is_saturation o1) x)^-1); ] + change in e1 with (∀x.b⎻* (A o2 x) = b'⎻* (A o2 x)); + apply (.= (e1 (a'⎻* (A o1 x)))); + alias symbol "invert" = "setoid1 symmetry". + lapply (†((saturated ?? a' (A o1 x) : ?) ^ -1)); + [2: apply (b'⎻* ); |4: apply Hletin; | skip; + |apply ((o_saturation_idempotent ?? (A_is_saturation o1) x)^-1);]] | intros; simplify; change with (((a34⎻* ∘ a23⎻* ) ∘ a12⎻* ) ∘ A o1 = ((a34⎻* ∘ (a23⎻* ∘ a12⎻* )) ∘ A o1)); - apply rule (#‡ASSOC1\sup -1); + apply rule (#‡ASSOC ^ -1); | intros; simplify; change with ((a⎻* ∘ (id2 ? o1)⎻* ) ∘ A o1 = a⎻* ∘ A o1); apply (#‡(id_neutral_right2 : ?)); @@ -181,6 +175,9 @@ definition BTop: category2. apply (#‡(id_neutral_left2 : ?));] qed. +definition btop_carr: BTop → Type1 ≝ λo:BTop. carr1 (oa_P (carrbt o)). +coercion btop_carr. + (* (*CSC: unused! *) (* this proof is more logic-oriented than set/lattice oriented *) @@ -206,4 +203,4 @@ theorem continuous_relation_eqS: [2,4: intros; apply saturation_monotone; try (apply A_is_saturation); apply Hcut;] apply Hcut2; assumption. qed. -*) \ No newline at end of file +*)