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=1081f5beab493cd652ed0872749ffc8f6a9f1cc3;hb=dba85ad7c7510e7bfd01e5721c63dad528b3e0bf;hp=c0fb6c6a7ff64f158bb667149ca1530ec497d69f;hpb=2d7053c212c790d528e82ba37c3e927070de7ae5;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 c0fb6c6a7..1081f5bea 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,44 +15,50 @@ include "o-algebra.ma". include "o-saturations.ma". -record basic_topology: Type ≝ +record basic_topology: Type2 ≝ { carrbt:> OA; - A: arrows1 SET (oa_P carrbt) (oa_P carrbt); - J: arrows1 SET (oa_P carrbt) (oa_P carrbt); + A: carrbt ⇒ carrbt; + J: carrbt ⇒ carrbt; A_is_saturation: is_saturation ? A; J_is_reduction: is_reduction ? J; compatibility: ∀U,V. (A U >< J V) = (U >< J V) }. -record continuous_relation (S,T: basic_topology) : Type ≝ - { cont_rel:> arrows1 ? S T; +lemma hint: OA → objs2 OA. + intro; apply t; +qed. +coercion hint. + +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); saturated: ∀U. U = A ? U → cont_rel⎻* U = A ? (cont_rel⎻* U) }. -definition continuous_relation_setoid: basic_topology → basic_topology → setoid1. +definition continuous_relation_setoid: basic_topology → basic_topology → setoid2. 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 ?)); - | simplify; intros; apply refl1; - | simplify; intros; apply sym1; apply H - | simplify; intros; apply trans1; [2: apply H |3: apply H1; |1: skip]]] + | 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 → arrows1 ? S T ≝ cont_rel. +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. - continuous_relation_setoid S T → unary_morphism (oa_P (carrbt S)) (oa_P (carrbt T)). + carr2 (continuous_relation_setoid S T) → ORelation_setoid (carrbt S) (carrbt T). intros; apply rule cont_rel; apply c; qed. coercion cont_rel''. + (* theorem continuous_relation_eq': ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2. @@ -101,68 +107,77 @@ theorem continuous_relation_eq_inv': split; apply Hcut; [2: assumption | intro; apply sym1; apply H] qed. *) + definition continuous_relation_comp: ∀o1,o2,o3. continuous_relation_setoid o1 o2 → continuous_relation_setoid o2 o3 → continuous_relation_setoid o1 o3. intros (o1 o2 o3 r s); constructor 1; - [ apply (s ∘ r) + [ apply (s ∘ r); | intros; apply sym1; change in match ((s ∘ r) U) with (s (r U)); - (*BAD*) unfold FunClass_1_OF_carr1; - apply (.= ((reduced : ?)\sup -1)); + (**) 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 refl1] | intros; - apply sym; + apply sym1; change in match ((s ∘ r)⎻* U) with (s⎻* (r⎻* U)); apply (.= (saturated : ?)\sup -1); - [ apply (.= (saturated : ?)); [ assumption | apply refl ] - | apply refl]] + [ apply (.= (saturated : ?)); [ assumption | apply refl1 ] + | apply refl1]] qed. -definition BTop: category1. +definition BTop: category2. constructor 1; [ apply basic_topology | apply continuous_relation_setoid | intro; constructor 1; - [ apply id1 - | intros; apply H; - | intros; apply H;] + [ apply id2 + | intros; apply e; + | intros; apply e;] | intros; constructor 1; [ apply continuous_relation_comp; - | intros; simplify; (*intro x; simplify;*) + | intros; simplify; + change with ((b⎻* ∘ a⎻* ) ∘ A o1 = ((b'⎻* ∘ a'⎻* ) ∘ A o1)); change with (b⎻* ∘ (a⎻* ∘ A o1) = b'⎻* ∘ (a'⎻* ∘ A o1)); - change in H with (a⎻* ∘ A o1 = a'⎻* ∘ A o1); - change in H1 with (b⎻* ∘ A o2 = b'⎻* ∘ A o2); - apply (.= H‡#); + 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 (eq1 (oa_P (carrbt 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 (oa_P (carrbt o2)) (a'⎻* (A o1 x)) (A o2 (a'⎻* (A o1 x)))); - unfold uncurry_arrows; - apply (.= †X); whd in H1; - lapply (H1 (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); - unfold uncurry_arrows; - apply (†(X\sup -1));] + change with (b⎻* (a'⎻* (A o1 x)) = b'⎻*(a'⎻* (A o1 x))); + alias symbol "trans" = "trans1". + alias symbol "prop1" = "prop11". + alias symbol "invert" = "setoid1 symmetry". + lapply (.= †(saturated o1 o2 a' (A o1 x) : ?)); + [3: apply (b⎻* ); | 5: apply Hletin; |1,2: skip; + |apply ((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 ((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⎻* ∘ (id1 ? o1)⎻* ) ∘ A o1 = a⎻* ∘ A o1); - apply (#‡(id_neutral_right1 : ?)); + change with ((a⎻* ∘ (id2 ? o1)⎻* ) ∘ A o1 = a⎻* ∘ A o1); + apply (#‡(id_neutral_right2 : ?)); | intros; simplify; - change with (((id1 ? o2)⎻* ∘ a⎻* ) ∘ A o1 = a⎻* ∘ A o1); - apply (#‡(id_neutral_left1 : ?));] + change with (((id2 ? o2)⎻* ∘ a⎻* ) ∘ A o1 = a⎻* ∘ A o1); + 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 *) @@ -188,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 +*)