X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_transition%2Flfpx_lfdeq.ma;h=668246dd1c56a09f551083e30d85e662336485ca;hb=5c186c72f508da0849058afeecc6877cd9ed6303;hp=3e77508453a306a50183d88c7942295145b98730;hpb=a09b60bd574adf1a7d3e423023009cb20c79d449;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma index 3e7750845..668246dd1 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma @@ -12,31 +12,36 @@ (* *) (**************************************************************************) -include "basic_2/relocation/lifts_tdeq.ma". -include "basic_2/static/lfxs_lfxs.ma". include "basic_2/static/lfdeq_fqup.ma". -include "basic_2/rt_transition/lfpx_frees.ma". -include "basic_2/rt_transition/lfpx.ma". (**) (* should be in lfpx_frees.ma *) +include "basic_2/static/lfdeq_lfdeq.ma". +include "basic_2/rt_transition/lfpx_fsle.ma". -(* UNCOUNTED PARALLEL RT-TRANSITION FOR LOCAL ENV.S ON REFERRED ENTRIES *****) +(* UNBOUND PARALLEL RT-TRANSITION FOR REFERRED LOCAL ENVIRONMENTS ***********) (* Properties with degree-based equivalence for local environments **********) lemma lfpx_pair_sn_split: ∀h,G,L1,L2,V. ⦃G, L1⦄ ⊢ ⬈[h, V] L2 → ∀o,I,T. - ∃∃L. ⦃G, L1⦄ ⊢ ⬈[h, ②{I}V.T] L & L ≡[h, o, V] L2. -/3 width=5 by lfpx_frees_conf, lfxs_pair_sn_split/ qed-. + ∃∃L. ⦃G, L1⦄ ⊢ ⬈[h, ②{I}V.T] L & L ≛[h, o, V] L2. +/3 width=5 by lfpx_fsge_comp, lfxs_pair_sn_split/ qed-. lemma lfpx_flat_dx_split: ∀h,G,L1,L2,T. ⦃G, L1⦄ ⊢ ⬈[h, T] L2 → ∀o,I,V. - ∃∃L. ⦃G, L1⦄ ⊢ ⬈[h, ⓕ{I}V.T] L & L ≡[h, o, T] L2. -/3 width=5 by lfpx_frees_conf, lfxs_flat_dx_split/ qed-. + ∃∃L. ⦃G, L1⦄ ⊢ ⬈[h, ⓕ{I}V.T] L & L ≛[h, o, T] L2. +/3 width=5 by lfpx_fsge_comp, lfxs_flat_dx_split/ qed-. lemma lfpx_bind_dx_split: ∀h,I,G,L1,L2,V1,T. ⦃G, L1.ⓑ{I}V1⦄ ⊢ ⬈[h, T] L2 → ∀o,p. - ∃∃L,V. ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V1.T] L & L.ⓑ{I}V ≡[h, o, T] L2 & ⦃G, L1⦄ ⊢ V1 ⬈[h] V. -/3 width=5 by lfpx_frees_conf, lfxs_bind_dx_split/ qed-. + ∃∃L,V. ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V1.T] L & L.ⓑ{I}V ≛[h, o, T] L2 & ⦃G, L1⦄ ⊢ V1 ⬈[h] V. +/3 width=5 by lfpx_fsge_comp, lfxs_bind_dx_split/ qed-. lemma lfpx_bind_dx_split_void: ∀h,G,K1,L2,T. ⦃G, K1.ⓧ⦄ ⊢ ⬈[h, T] L2 → ∀o,p,I,V. - ∃∃K2. ⦃G, K1⦄ ⊢ ⬈[h, ⓑ{p,I}V.T] K2 & K2.ⓧ ≡[h, o, T] L2. -/3 width=5 by lfpx_frees_conf, lfxs_bind_dx_split_void/ qed-. + ∃∃K2. ⦃G, K1⦄ ⊢ ⬈[h, ⓑ{p,I}V.T] K2 & K2.ⓧ ≛[h, o, T] L2. +/3 width=5 by lfpx_fsge_comp, lfxs_bind_dx_split_void/ qed-. + +lemma lfpx_tdeq_conf: ∀h,o,G. s_r_confluent1 … (cdeq h o) (lfpx h G). +/2 width=5 by tdeq_lfxs_conf/ qed-. + +lemma lfpx_tdeq_div: ∀h,o,T1,T2. T1 ≛[h, o] T2 → + ∀G,L1,L2. ⦃G, L1⦄ ⊢ ⬈[h, T2] L2 → ⦃G, L1⦄ ⊢ ⬈[h, T1] L2. +/2 width=5 by tdeq_lfxs_div/ qed-. lemma cpx_tdeq_conf_lexs: ∀h,o,G. R_confluent2_lfxs … (cpx h G) (cdeq h o) (cpx h G) (cdeq h o). #h #o #G #L0 #T0 #T1 #H @(cpx_ind … H) -G -L0 -T0 -T1 /2 width=3 by ex2_intro/ @@ -119,45 +124,44 @@ lemma cpx_tdeq_conf_lexs: ∀h,o,G. R_confluent2_lfxs … (cpx h G) (cdeq h o) ( qed-. lemma cpx_tdeq_conf: ∀h,o,G,L. ∀T0:term. ∀T1. ⦃G, L⦄ ⊢ T0 ⬈[h] T1 → - ∀T2. T0 ≡[h, o] T2 → - ∃∃T. T1 ≡[h, o] T & ⦃G, L⦄ ⊢ T2 ⬈[h] T. + ∀T2. T0 ≛[h, o] T2 → + ∃∃T. T1 ≛[h, o] T & ⦃G, L⦄ ⊢ T2 ⬈[h] T. #h #o #G #L #T0 #T1 #HT01 #T2 #HT02 elim (cpx_tdeq_conf_lexs … HT01 … HT02 L … L) -HT01 -HT02 /2 width=3 by lfxs_refl, ex2_intro/ qed-. -lemma tdeq_cpx_trans: ∀h,o,G,L,T2. ∀T0:term. T2 ≡[h, o] T0 → +lemma tdeq_cpx_trans: ∀h,o,G,L,T2. ∀T0:term. T2 ≛[h, o] T0 → ∀T1. ⦃G, L⦄ ⊢ T0 ⬈[h] T1 → - ∃∃T. ⦃G, L⦄ ⊢ T2 ⬈[h] T & T ≡[h, o] T1. + ∃∃T. ⦃G, L⦄ ⊢ T2 ⬈[h] T & T ≛[h, o] T1. #h #o #G #L #T2 #T0 #HT20 #T1 #HT01 elim (cpx_tdeq_conf … HT01 T2) -HT01 /3 width=3 by tdeq_sym, ex2_intro/ qed-. (* Basic_2A1: uses: cpx_lleq_conf *) lemma cpx_lfdeq_conf: ∀h,o,G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ⬈[h] T1 → - ∀L2. L0 ≡[h, o, T0] L2 → - ∃∃T. ⦃G, L2⦄ ⊢ T0 ⬈[h] T & T1 ≡[h, o] T. + ∀L2. L0 ≛[h, o, T0] L2 → + ∃∃T. ⦃G, L2⦄ ⊢ T0 ⬈[h] T & T1 ≛[h, o] T. #h #o #G #L0 #T0 #T1 #HT01 #L2 #HL02 elim (cpx_tdeq_conf_lexs … HT01 T0 … L0 … HL02) -HT01 -HL02 /2 width=3 by lfxs_refl, ex2_intro/ qed-. (* Basic_2A1: uses: lleq_cpx_trans *) -lemma lfdeq_cpx_trans: ∀h,o,G,L2,L0,T0. L2 ≡[h, o, T0] L0 → +lemma lfdeq_cpx_trans: ∀h,o,G,L2,L0,T0. L2 ≛[h, o, T0] L0 → ∀T1. ⦃G, L0⦄ ⊢ T0 ⬈[h] T1 → - ∃∃T. ⦃G, L2⦄ ⊢ T0 ⬈[h] T & T ≡[h, o] T1. + ∃∃T. ⦃G, L2⦄ ⊢ T0 ⬈[h] T & T ≛[h, o] T1. #h #o #G #L2 #L0 #T0 #HL20 #T1 #HT01 elim (cpx_lfdeq_conf … o … HT01 L2) -HT01 /3 width=3 by lfdeq_sym, tdeq_sym, ex2_intro/ qed-. lemma lfpx_lfdeq_conf: ∀h,o,G,T. confluent2 … (lfpx h G T) (lfdeq h o T). -/3 width=6 by lfpx_frees_conf, cpx_tdeq_conf_lexs, frees_lfdeq_conf_lexs, lfxs_conf/ qed-. +/3 width=6 by lfpx_fsge_comp, lfdeq_fsge_comp, cpx_tdeq_conf_lexs, lfxs_conf/ qed-. -(* Basic_2A1: uses: lleq_lpx_trans *) lemma lfdeq_lfpx_trans: ∀h,o,G,T,L2,K2. ⦃G, L2⦄ ⊢ ⬈[h, T] K2 → - ∀L1. L1 ≡[h, o, T] L2 → - ∃∃K1. ⦃G, L1⦄ ⊢ ⬈[h, T] K1 & K1 ≡[h, o, T] K2. + ∀L1. L1 ≛[h, o, T] L2 → + ∃∃K1. ⦃G, L1⦄ ⊢ ⬈[h, T] K1 & K1 ≛[h, o, T] K2. #h #o #G #T #L2 #K2 #HLK2 #L1 #HL12 elim (lfpx_lfdeq_conf … o … HLK2 L1) /3 width=3 by lfdeq_sym, ex2_intro/