From: Claudio Sacerdoti Coen Date: Sat, 3 Jan 2009 16:41:15 +0000 (+0000) Subject: More re-working. X-Git-Tag: make_still_working~4307 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=06585b97fad3158391dbbea1fcad5866f5269eee;p=helm.git More re-working. --- diff --git a/helm/software/matita/contribs/formal_topology/overlap/o-algebra.ma b/helm/software/matita/contribs/formal_topology/overlap/o-algebra.ma index 029e93258..533bff3dd 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/o-algebra.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/o-algebra.ma @@ -335,3 +335,13 @@ lemma setoid1_of_OA: OA → setoid1. intro; apply (oa_P t); qed. coercion setoid1_of_OA. + +lemma SET1_of_OA: OA → SET1. + intro; whd; apply (setoid1_of_OA t); +qed. +coercion SET1_of_OA. + +lemma objs2_SET1_OF_OA: OA → objs2 SET1. + intro; whd; apply (setoid1_of_OA t); +qed. +coercion objs2_SET1_OF_OA. 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..64c867e22 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 @@ -17,31 +17,37 @@ include "o-saturations.ma". record basic_topology: Type ≝ { 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) }. +lemma hint: OA → objs2 OA. + intro; apply t; +qed. +coercion hint. + record continuous_relation (S,T: basic_topology) : Type ≝ - { cont_rel:> arrows1 ? S T; + { 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. coercion cont_rel'. @@ -53,6 +59,8 @@ definition cont_rel'': qed. coercion cont_rel''. +*) + (* theorem continuous_relation_eq': ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2.