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 "basic_2/reduction/cpr.ma".
17 (* EXAMPLES *****************************************************************)
19 (* A reduction cycle in two steps: the term Omega ***************************)
21 definition Delta: term → term ≝ λW. +ⓛW.ⓐ#0.#0.
23 definition Omega1: term → term ≝ λW. ⓐ(Delta W).(Delta W).
25 definition Omega2: term → term ≝ λW. +ⓓⓝW.(Delta W).ⓐ#0.#0.
27 (* Basic properties *********************************************************)
29 lemma Delta_lift: ∀W1,W2,d,e. ⇧[d, e] W1 ≡ W2 →
30 ⇧[d, e] (Delta W1) ≡ (Delta W2).
31 /4 width=1 by lift_flat, lift_bind, lift_lref_lt/ qed.
33 (* Main properties **********************************************************)
35 theorem cpr_Omega_12: ∀W. ⦃⋆, ⋆⦄ ⊢ Omega1 W ➡ Omega2 W.
36 /2 width=1 by cpr_beta/ qed.
38 theorem cpr_Omega_21: ∀W. ⦃⋆, ⋆⦄ ⊢ Omega2 W ➡ Omega1 W.
39 #W1 elim (lift_total W1 0 1) #W2 #HW12
40 @(cpr_zeta … (Omega1 W2)) /3 width=1 by Delta_lift, lift_flat/
41 @cpr_flat @(cpr_delta … (Delta W1) ? 0)
42 [3,5,8,10: /2 width=2 by Delta_lift/ |4,9: /2 width=1 by cpr_eps/ |*: skip ]