X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fformal_topology%2Foverlap%2Fsaturations.ma;h=cc0db526c340daecc942529fce614cc12edfa433;hb=05cfeb82d2624860e66941421a937f308d66cf33;hp=b78952fdb4c6f2a0a50f394e1c47990a2742f751;hpb=5be81fce195f2b45ec57c5422d35e4c03827891d;p=helm.git diff --git a/helm/software/matita/contribs/formal_topology/overlap/saturations.ma b/helm/software/matita/contribs/formal_topology/overlap/saturations.ma index b78952fdb..cc0db526c 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/saturations.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/saturations.ma @@ -14,13 +14,11 @@ include "relations.ma". -definition is_saturation: ∀C:REL. unary_morphism1 (Ω \sup C) (Ω \sup C) → CProp1 ≝ - λC:REL.λA:unary_morphism1 (Ω \sup C) (Ω \sup C). - ∀U,V. (U ⊆ A V) = (A U ⊆ A V). +definition is_saturation: ∀C:REL. Ω^C ⇒_1 Ω^C → CProp1 ≝ + λC:REL.λA:Ω^C ⇒_1 Ω^C. ∀U,V. (U ⊆ A V) =_1 (A U ⊆ A V). -definition is_reduction: ∀C:REL. unary_morphism1 (Ω \sup C) (Ω \sup C) → CProp1 ≝ - λC:REL.λJ:unary_morphism1 (Ω \sup C) (Ω \sup C). - ∀U,V. (J U ⊆ V) = (J U ⊆ J V). +definition is_reduction: ∀C:REL. Ω^C ⇒_1 Ω^C → CProp1 ≝ + λC:REL.λJ: Ω^C ⇒_1 Ω^C. ∀U,V. (J U ⊆ V) =_1 (J U ⊆ J V). theorem saturation_expansive: ∀C,A. is_saturation C A → ∀U. U ⊆ A U. intros; apply (fi ?? (i ??)); apply subseteq_refl.