1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "datatypes/subsets.ma".
16 include "formal_topology/relations.ma".
18 definition is_saturation ≝
19 λC:REL.λA:unary_morphism (Ω \sup C) (Ω \sup C).
20 ∀U,V. (U ⊆ A V) = (A U ⊆ A V).
22 definition is_reduction ≝
23 λC:REL.λJ:unary_morphism (Ω \sup C) (Ω \sup C).
24 ∀U,V. (J U ⊆ V) = (J U ⊆ J V).
26 theorem saturation_expansive: ∀C,A. is_saturation C A → ∀U. U ⊆ A U.
27 intros; apply (fi ?? (H ??)); apply subseteq_refl.
30 theorem saturation_monotone:
31 ∀C,A. is_saturation C A →
32 ∀U,V. U ⊆ V → A U ⊆ A V.
33 intros; apply (if ?? (H ??)); apply subseteq_trans; [apply V|3: apply saturation_expansive ]
37 theorem saturation_idempotent: ∀C,A. is_saturation C A → ∀U. A (A U) = A U.
39 [ apply (if ?? (H ??)); apply subseteq_refl
40 | apply saturation_expansive; assumption]