X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fapps_2%2Fexamples%2Fex_fpbg_refl.ma;h=f15bbd19c1bde336d8f8b33c8a60e2d1352ba4ed;hp=e37f8831f0689d63dfb49c72c875ad3da8ab7fd0;hb=d71e53021b0c17e1a00c2d623e7139c6d18069d5;hpb=f9abd21eb0d26cf9b632af4df819225be4d091e3 diff --git a/matita/matita/contribs/lambdadelta/apps_2/examples/ex_fpbg_refl.ma b/matita/matita/contribs/lambdadelta/apps_2/examples/ex_fpbg_refl.ma index e37f8831f..f15bbd19c 100644 --- a/matita/matita/contribs/lambdadelta/apps_2/examples/ex_fpbg_refl.ma +++ b/matita/matita/contribs/lambdadelta/apps_2/examples/ex_fpbg_refl.ma @@ -12,11 +12,13 @@ (* *) (**************************************************************************) -include "basic_2/computation/fpbg_fpbs.ma". +include "basic_2/rt_computation/fpbs_cpxs.ma". +include "basic_2/rt_computation/fpbg_fqup.ma". +include "basic_2/rt_computation/fpbg_fpbs.ma". (* EXAMPLES *****************************************************************) -(* Reflexivity of proper qrst-computation: the term ApplOmega ***************) +(* Reflexivity of proper rst-computation: the term ApplOmega ****************) definition ApplDelta: term → nat → term ≝ λW,s. +ⓛW.ⓐ⋆s.ⓐ#0.#0. @@ -28,27 +30,28 @@ definition ApplOmega3: term → nat → term ≝ λW,s. ⓐ⋆s.(ApplOmega1 W s) (* Basic properties *********************************************************) -lemma ApplDelta_lift: ∀W1,W2,s,l,k. ⬆[l, k] W1 ≡ W2 → - ⬆[l, k] (ApplDelta W1 s) ≡ (ApplDelta W2 s). -/5 width=1 by lift_flat, lift_bind, lift_lref_lt/ qed. +lemma ApplDelta_lifts (f:rtmap): + ∀W1,W2,s. ⬆*[f] W1 ≘ W2 → + ⬆*[f] (ApplDelta W1 s) ≘ (ApplDelta W2 s). +/5 width=1 by lifts_sort, lifts_lref, lifts_bind, lifts_flat/ qed. -lemma cpr_ApplOmega_12: ∀G,L,W,s. ⦃G, L⦄ ⊢ ApplOmega1 W s ➡ ApplOmega2 W s. -/2 width=1 by cpr_beta/ qed. +lemma cpr_ApplOmega_12 (h): ∀G,L,W,s. ⦃G, L⦄ ⊢ ApplOmega1 W s ➡[h] ApplOmega2 W s. +/2 width=1 by cpm_beta/ qed. -lemma cpr_ApplOmega_23: ∀G,L,W,s. ⦃G, L⦄ ⊢ ApplOmega2 W s ➡ ApplOmega3 W s. -#G #L #W1 #s elim (lift_total W1 0 1) #W2 #HW12 -@(cpr_zeta … (ApplOmega3 W2 s)) /4 width=1 by ApplDelta_lift, lift_flat/ -@cpr_flat // @cpr_flat @(cpr_delta … (ApplDelta W1 s) ? 0) -[3,5,8,10: /2 width=2 by ApplDelta_lift/ |4,9: /2 width=1 by cpr_eps/ |*: skip ] +lemma cpr_ApplOmega_23 (h): ∀G,L,W,s. ⦃G, L⦄ ⊢ ApplOmega2 W s ➡[h] ApplOmega3 W s. +#h #G #L #W1 #s elim (lifts_total W1 (𝐔❴1❵)) #W2 #HW12 +@(cpm_zeta … (ApplOmega3 W2 s)) /4 width=1 by ApplDelta_lifts, lifts_flat/ +@cpm_appl // @cpm_appl @(cpm_delta … (ApplDelta W1 s)) +/2 width=1 by ApplDelta_lifts, cpm_eps/ qed. -lemma cpxs_ApplOmega_13: ∀h,o,G,L,W,s. ⦃G, L⦄ ⊢ ApplOmega1 W s ➡*[h,o] ApplOmega3 W s. -/4 width=3 by cpxs_strap1, cpr_cpx/ qed. +lemma cpxs_ApplOmega_13 (h): ∀G,L,W,s. ⦃G, L⦄ ⊢ ApplOmega1 W s ⬈*[h] ApplOmega3 W s. +/4 width=4 by cpxs_strap1, cpm_fwd_cpx/ qed. lemma fqup_ApplOmega_13: ∀G,L,W,s. ⦃G, L, ApplOmega3 W s⦄ ⊐+ ⦃G, L, ApplOmega1 W s⦄. /2 width=1 by/ qed. (* Main properties **********************************************************) -theorem fpbg_refl: ∀h,o,G,L,W,s. ⦃G, L, ApplOmega1 W s⦄ >≡[h,o] ⦃G, L, ApplOmega1 W s⦄. +theorem fpbg_refl (h) (o): ∀G,L,W,s. ⦃G, L, ApplOmega1 W s⦄ >[h,o] ⦃G, L, ApplOmega1 W s⦄. /3 width=5 by fpbs_fpbg_trans, fqup_fpbg, cpxs_fpbs/ qed.