From 6ed4f0127b8acb6caeba6fbfadef7f990dd7803e Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Mon, 27 Nov 2017 20:06:19 +0000 Subject: [PATCH] work in progress ... --- .../basic_2/rt_computation/lfpxs_etc.ma | 24 ++--- .../basic_2/rt_transition/cpx_etc.ma | 93 +++++++++++++++++++ .../lambdadelta/basic_2/static/lfxs_fle.ma | 1 - 3 files changed, 102 insertions(+), 16 deletions(-) create mode 100644 matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_etc.ma diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_etc.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_etc.ma index 0fda037b1..05f1e7357 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_etc.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_etc.ma @@ -1,20 +1,14 @@ -include "basic_2/static/lfxs_lfxs.ma". -include "basic_2/rt_transition/lfpx_frees.ma". -include "basic_2/rt_computation/lfpxs_fqup.ma". - -axiom cpx_frees_conf_lfpxs: ∀h,G,L1,T1,T2. ⦃G, L1⦄ ⊢ T1 ⬈[h] T2 → - ∀f1. L1 ⊢ 𝐅*⦃T1⦄ ≡ f1 → - ∀L2. ⦃G, L1⦄ ⊢ ⬈*[h, T1] L2 → - ∀g1. L2 ⊢ 𝐅*⦃T1⦄ ≡ g1 → - ∃∃g2. L2 ⊢ 𝐅*⦃T2⦄ ≡ g2 & g2 ⊆ g1 & g1 ⊆ f1. +include "basic_2/static/lfxs_lex.ma". +include "basic_2/static/lfxs_fle.ma". +include "basic_2/rt_transition/cpx_etc.ma". +include "basic_2/rt_computation/lfpxs_lpxs.ma". lemma lfpxs_cpx_conf: ∀h,G. s_r_confluent1 … (cpx h G) (lfpxs h G). #h #G #L1 #T1 #T2 #HT12 #L2 #H -lapply (cpx_frees_conf_lfpxs … HT12) -HT12 #HT12 -@(lfpxs_ind_sn … H) -L2 // -#L #L2 #HL1 * #g1 #Hg1 #HL2 #IH -elim (frees_total L1 T1) #f1 #Hf1 -elim (HT12 … Hf1 … HL1 … Hg1) -T1 #g2 #Hg2 #Hg21 #_ -f1 -/4 width=7 by lfpxs_step_dx, sle_lexs_trans, ex2_intro/ +elim (tc_lfxs_inv_lex_lfeq … H) -H #L #HL1 #HL2 +lapply (lfxs_lex … HL1 T1) #H +elim (cpx_lfxs_conf_gen … HT12 … H) -HT12 -H // #_ #HT21 #_ +@(lfpxs_lpxs_lfeq … HL1) -HL1 +@(fle_lfxs_trans … HL2) // qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_etc.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_etc.ma new file mode 100644 index 000000000..7a516ee46 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_etc.ma @@ -0,0 +1,93 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/fle_drops.ma". +include "basic_2/static/fle_fqup.ma". +include "basic_2/static/fle_fle.ma". +include "basic_2/static/lfxs.ma". +include "basic_2/rt_transition/cpx.ma". + +axiom fle_pair_sn: ∀L,V1,V2. ⦃L, V1⦄ ⊆ ⦃L, V2⦄ → + ∀I1,I2,T. ⦃L.ⓑ{I1}V1, T⦄ ⊆ ⦃L.ⓑ{I2}V2, T⦄. +(* +axiom fle_elim_flat: ∀L1,L2,V1,T2. ⦃L1, V1⦄ ⊆ ⦃L2, T2⦄ → ∀T1. ⦃L1, T1⦄ ⊆ ⦃L2, T2⦄ → + ∀I. ⦃L1, ⓕ{I}V1.T1⦄ ⊆ ⦃L2, T2⦄. + +axiom fle_elim_bind: + ∀p,I1,I2,J1,J2,L,V,W,T. ⦃L, ⓑ{p,I1}ⓕ{J1}V.W.T⦄ ⊆ ⦃L, ⓕ{J2}V.ⓑ{p,I2}W.T⦄. +*) +(* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS *************) + +(* Properties with context-sensitive free variables *************************) + +(* Note: "⦃L2, T1⦄ ⊆ ⦃L0, T1⦄" may not hold *) +axiom cpx_lfxs_conf_gen: ∀R. c_reflexive … R → + ∀h,G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ⬈[h] T1 → + ∀L2. L0 ⪤*[R, T0] L2 → + ∧∧ ⦃L2, T0⦄ ⊆ ⦃L0, T0⦄ & ⦃L2, T1⦄ ⊆ ⦃L2, T0⦄ + & ⦃L0, T1⦄ ⊆ ⦃L0, T0⦄. +(* +#R #HR #h #G #L0 #T0 @(fqup_wf_ind_eq (Ⓕ) … G L0 T0) -G -L0 -T0 +#G #L #T #IH #G0 #L0 * * +[ #s #HG #HL #HT #X #HX #Y #_ destruct -IH + elim (cpx_inv_sort1 … HX) -HX #H destruct + /2 width=1 by and3_intro/ +| +| #l #HG #HL #HT #X #HX #Y #_ destruct -IH + >(cpx_inv_gref1 … HX) -X + /2 width=1 by and3_intro/ +| #p #I #V0 #T0 #HG #HL #HT #X #HX #Y #HY destruct + elim (lfxs_inv_bind … V0 ? HY) -HY // #HV0 #HT0 + elim (cpx_inv_bind1 … HX) -HX * + [ #V1 #T1 #HV01 #HT01 #H destruct + elim (IH … HV01 … HV0) -HV01 -HV0 // #H1V #H2V #H3V + elim (IH … HT01 … HT0) -HT01 -HT0 -IH // #H1T #H2T #H3T + /5 width=4 by fle_pair_sn, fle_trans, fle_bind, and3_intro/ + | #T #HT #HXT #H1 #H2 destruct + elim (IH G0 … V0… V0 … HV0) -HV0 // #H1V #H2V #H3V + elim (IH … HT … HT0) -HT -HT0 -IH // #H1T #H2T #H3T + /4 width=7 by fle_bind_dx, fle_trans, fle_bind, and3_intro/ + ] +| #I #V0 #X0 #HG #HL #HT #X #HX #Y #HY destruct + elim (lfxs_inv_flat … HY) -HY #HV0 #HX0 + elim (cpx_inv_flat1 … HX) -HX * + [ #V1 #T1 #HV01 #HT01 #H destruct + elim (IH … HV01 … HV0) -HV01 -HV0 // #H1V #H2V #H3V + elim (IH … HT01 … HX0) -HT01 -HX0 -IH // #H1T #H2T #H3T + /3 width=3 by fle_flat, and3_intro/ + | #HX #H destruct + elim (IH G0 … V0… V0 … HV0) -HV0 // #H1V #H2V #H3V + elim (IH … HX … HX0) -HX -HX0 -IH // #H1T #H2T #H3T + /4 width=4 by fle_trans, fle_flat, and3_intro/ + | #HX #H destruct + elim (IH … HX … HV0) -HX -HV0 // #H1V #H2V #H3V + elim (IH G0 … X0… X0 … HX0) -HX0 -IH // #H1T #H2T #H3T + /4 width=4 by fle_trans, fle_flat, and3_intro/ + | #p #V1 #W0 #W1 #T0 #T1 #HV01 #HW01 #HT01 #H1 #H2 #H3 destruct + elim (lfxs_inv_bind … W0 ? HX0) -HX0 // #HW0 #HT0 + elim (IH … HV01 … HV0) -HV01 -HV0 // #H1V #H2V #H3V + elim (IH … HW01 … HW0) -HW01 -HW0 // #H1W #H2W #H3W + elim (IH … HT01 … HT0) -HT01 -HT0 -IH // #H1T #H2T #H3T + @and3_intro + [ /3 width=5 by fle_flat, fle_bind/ + | @(fle_trans) + [4: @(fle_flat … H2V) [2: @(fle_bind … H2T) // | skip | @Appl ] + |1,2: skip + | @(fle_trans) [3: @fle_elim_bind + + … H2T) // + + // |1,2: skip + | /2 width=1 by fle_bind_dx/ +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_fle.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_fle.ma index df8aec011..0d851ce6a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_fle.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_fle.ma @@ -12,7 +12,6 @@ (* *) (**************************************************************************) -include "basic_2/static/frees_frees.ma". include "basic_2/static/fle.ma". include "basic_2/static/lfxs_lfxs.ma". -- 2.39.2