X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fformal_topology%2Fsaturations.ma;fp=matita%2Fmatita%2Flib%2Fformal_topology%2Fsaturations.ma;h=02030e7cde3f2a2789a88f7f936e7bfc3ff98732;hb=c8718cc46ab9aaca047366dfefe72bc7c9402e5a;hp=0000000000000000000000000000000000000000;hpb=000dc5a8de79b2ab63a49cf0f9db2b540cc05bcf;p=helm.git diff --git a/matita/matita/lib/formal_topology/saturations.ma b/matita/matita/lib/formal_topology/saturations.ma new file mode 100644 index 000000000..02030e7cd --- /dev/null +++ b/matita/matita/lib/formal_topology/saturations.ma @@ -0,0 +1,38 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "formal_topology/relations.ma". + +definition is_saturation: ∀C:REL. Ω^C ⇒_1 Ω^C → CProp[1] ≝ + λC:REL.λA:Ω^C ⇒_1 Ω^C. ∀U,V. (U ⊆ A V) =_1 (A U ⊆ A V). + +definition is_reduction: ∀C:REL. Ω^C ⇒_1 Ω^C → CProp[1] ≝ + λ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. +qed. + +theorem saturation_monotone: + ∀C,A. is_saturation C A → + ∀U,V. U ⊆ V → A U ⊆ A V. + intros; apply (if ?? (i ??)); apply subseteq_trans; [apply V|3: apply saturation_expansive ] + assumption. +qed. + +theorem saturation_idempotent: ∀C,A. is_saturation C A → ∀U. A (A U) = A U. + intros; split; + [ apply (if ?? (i ??)); apply subseteq_refl + | apply saturation_expansive; assumption] +qed.