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=30a8b476cb2045f1ea204e4489dc34c1e54cc2e5;hb=a484c51de8ba6c56f02f9c0758688d3c9186b63d;hp=640a936ee21a194e1cd35a291ebce95dd2870d15;hpb=f1977440576f1385c0987e635dedbd874ad92d9c;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 640a936ee..30a8b476c 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 @@ -1,4 +1,4 @@ -(**************************************************************************) + (**************************************************************************) (* ___ *) (* ||M|| *) (* ||A|| A project by Andrea Asperti *) @@ -15,21 +15,16 @@ 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) }. -lemma hint: OA → objs2 OA. - intro; apply t; -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); @@ -40,24 +35,17 @@ definition continuous_relation_setoid: basic_topology → basic_topology → set intros (S T); constructor 1; [ apply (continuous_relation S T) | constructor 1; - [ (*apply (λr,s:continuous_relation S T.∀b. eq1 (oa_P (carrbt S)) (A ? (r⎻ b)) (A ? (s⎻ b)));*) - apply (λr,s:continuous_relation S T.r⎻* ∘ (A S) = s⎻* ∘ (A ?)); + [ alias symbol "eq" = "setoid2 eq". + alias symbol "compose" = "category2 composition". + apply (λr,s:continuous_relation S T. (r⎻* ) ∘ (A S) = (s⎻* ∘ (A ?))); | simplify; intros; apply refl2; | simplify; intros; apply sym2; apply e | simplify; intros; apply trans2; [2: apply e |3: apply e1; |1: skip]]] qed. -definition cont_rel': ∀S,T: basic_topology. continuous_relation_setoid S T → arrows2 ? S T ≝ cont_rel. - -coercion cont_rel'. - -definition cont_rel'': - ∀S,T: basic_topology. - carr2 (continuous_relation_setoid S T) → ORelation_setoid (carrbt S) (carrbt T). - intros; apply rule cont_rel; apply c; -qed. - -coercion cont_rel''. +definition continuous_relation_of_continuous_relation_setoid: + ∀P,Q. continuous_relation_setoid P Q → continuous_relation P Q ≝ λP,Q,c.c. +coercion continuous_relation_of_continuous_relation_setoid. (* theorem continuous_relation_eq': @@ -108,7 +96,6 @@ theorem continuous_relation_eq_inv': qed. *) -axiom daemon: False. definition continuous_relation_comp: ∀o1,o2,o3. continuous_relation_setoid o1 o2 → @@ -117,15 +104,10 @@ definition continuous_relation_comp: intros (o1 o2 o3 r s); constructor 1; [ apply (s ∘ r); | intros; - apply sym1; + apply sym1; change in match ((s ∘ r) U) with (s (r U)); - (**) unfold FunClass_1_OF_Type_OF_setoid2; - unfold objs2_OF_basic_topology1; unfold hint; - letin reduced := reduced; clearbody reduced; - unfold uncurry_arrows in reduced ⊢ %; (**) apply (.= (reduced : ?)\sup -1); - [ (*BAD*) change with (eq1 ? (r U) (J ? (r U))); - (* BAD U *) apply (.= (reduced ??? U ?)); [ assumption | apply refl1 ] + [ apply (.= (reduced :?)); [ assumption | apply refl1 ] | apply refl1] | intros; apply sym1; @@ -135,11 +117,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 @@ -156,20 +133,18 @@ definition BTop: category2. change in e with (a⎻* ∘ A o1 = a'⎻* ∘ A o1); change in e1 with (b⎻* ∘ A o2 = b'⎻* ∘ A o2); apply (.= e‡#); - intro x; - change with (b⎻* (a'⎻* (A o1 x)) = b'⎻*(a'⎻* (A o1 x))); - alias symbol "trans" = "trans1". - alias symbol "prop1" = "prop11". - apply (.= †(saturated o1 o2 a' (A o1 x) : ?)); - [ apply ((saturation_idempotent ?? (A_is_saturation o1) x)^-1); ] - change in e1 with (∀x.b⎻* (A o2 x) = b'⎻* (A o2 x)); + intro x; + change with (eq1 ? (b⎻* (a'⎻* (A o1 x))) (b'⎻*(a'⎻* (A o1 x)))); + apply (.= †(saturated o1 o2 a' (A o1 x) ?)); [ + apply ((o_saturation_idempotent ?? (A_is_saturation o1) x)^-1);] apply (.= (e1 (a'⎻* (A o1 x)))); - alias symbol "invert" = "setoid1 symmetry". - apply (†((saturated ?? a' (A o1 x) : ?) ^ -1)); - apply ((saturation_idempotent ?? (A_is_saturation o1) x)^-1);] + change with (eq1 ? (b'⎻* (A o2 (a'⎻* (A o1 x)))) (b'⎻*(a'⎻* (A o1 x)))); + apply (.= †(saturated o1 o2 a' (A o1 x):?)^-1); [ + apply ((o_saturation_idempotent ?? (A_is_saturation o1) x)^-1);] + apply rule #;] | 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 : ?)); @@ -178,6 +153,13 @@ definition BTop: category2. apply (#‡(id_neutral_left2 : ?));] qed. +definition basic_topology_of_BTop: objs2 BTop → basic_topology ≝ λx.x. +coercion basic_topology_of_BTop. + +definition continuous_relation_setoid_of_arrows2_BTop : + ∀P,Q. arrows2 BTop P Q → continuous_relation_setoid P Q ≝ λP,Q,x.x. +coercion continuous_relation_setoid_of_arrows2_BTop. + (* (*CSC: unused! *) (* this proof is more logic-oriented than set/lattice oriented *) @@ -203,4 +185,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 +*)