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 "relations.ma".
17 definition is_saturation: ∀C:REL. unary_morphism1 (Ω \sup C) (Ω \sup C) → CProp1 ≝
18 λC:REL.λA:unary_morphism1 (Ω \sup C) (Ω \sup C).
19 ∀U,V. (U ⊆ A V) = (A U ⊆ A V).
21 definition is_reduction: ∀C:REL. unary_morphism1 (Ω \sup C) (Ω \sup C) → CProp1 ≝
22 λC:REL.λJ:unary_morphism1 (Ω \sup C) (Ω \sup C).
23 ∀U,V. (J U ⊆ V) = (J U ⊆ J V).
25 theorem saturation_expansive: ∀C,A. is_saturation C A → ∀U. U ⊆ A U.
26 intros; apply (fi ?? (i ??)); apply subseteq_refl.
29 theorem saturation_monotone:
30 ∀C,A. is_saturation C A →
31 ∀U,V. U ⊆ V → A U ⊆ A V.
32 intros; apply (if ?? (i ??)); apply subseteq_trans; [apply V|3: apply saturation_expansive ]
36 theorem saturation_idempotent: ∀C,A. is_saturation C A → ∀U. A (A U) = A U.
38 [ apply (if ?? (i ??)); apply subseteq_refl
39 | apply saturation_expansive; assumption]