From b6e1db4f1b0f1d5121f2b214562f96c5b0fa544e Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Mon, 4 Jul 2016 21:06:25 +0000 Subject: [PATCH] basic properties of cpr ... --- .../lambdadelta/basic_2/rt_transition/cpm.ma | 18 + .../lambdadelta/basic_2/rt_transition/cpr.etc | 87 +++++ .../lambdadelta/basic_2/rt_transition/cpr.ma | 325 +++--------------- .../basic_2/rt_transition/partial.txt | 1 + .../lambdadelta/ground_2/lib/arith.ma | 4 + .../lambdadelta/ground_2/steps/rtc_max.ma | 84 +++++ .../lambdadelta/ground_2/steps/rtc_plus.ma | 9 +- 7 files changed, 252 insertions(+), 276 deletions(-) create mode 100644 matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpr.etc create mode 100644 matita/matita/contribs/lambdadelta/ground_2/steps/rtc_max.ma diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm.ma index 3201a5b8f..7138b3210 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm.ma @@ -53,6 +53,7 @@ lemma cpm_lref: ∀n,h,I,G,K,V,T,U,i. ⦃G, K⦄ ⊢ #i ➡[n, h] T → /3 width=5 by cpg_lref, ex2_intro/ qed. +(* Basic_2A1: includes: cpr_bind *) lemma cpm_bind: ∀n,h,p,I,G,L,V1,V2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V2 → ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡[n, h] T2 → ⦃G, L⦄ ⊢ ⓑ{p,I}V1.T1 ➡[n, h] ⓑ{p,I}V2.T2. @@ -60,6 +61,9 @@ lemma cpm_bind: ∀n,h,p,I,G,L,V1,V2,T1,T2. /5 width=5 by cpg_bind, isrt_plus_O1, isr_shift, ex2_intro/ qed. +(* Note: cpr_flat: does not hold in basic_1 *) +(* Basic_1: includes: pr2_thin_dx *) +(* Basic_2A1: includes: cpr_flat *) lemma cpm_flat: ∀n,h,I,G,L,V1,V2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V2 → ⦃G, L⦄ ⊢ T1 ➡[n, h] T2 → ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ➡[n, h] ⓕ{I}V2.T2. @@ -67,12 +71,14 @@ lemma cpm_flat: ∀n,h,I,G,L,V1,V2,T1,T2. /5 width=5 by isrt_plus_O1, isr_shift, cpg_flat, ex2_intro/ qed. +(* Basic_2A1: includes: cpr_zeta *) lemma cpm_zeta: ∀n,h,G,L,V,T1,T,T2. ⦃G, L.ⓓV⦄ ⊢ T1 ➡[n, h] T → ⬆*[1] T2 ≡ T → ⦃G, L⦄ ⊢ +ⓓV.T1 ➡[n, h] T2. #n #h #G #L #V #T1 #T #T2 * /3 width=5 by cpg_zeta, isrt_plus_O2, ex2_intro/ qed. +(* Basic_2A1: includes: cpr_eps *) lemma cpm_eps: ∀n,h,G,L,V,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[n, h] T2 → ⦃G, L⦄ ⊢ ⓝV.T1 ➡[n, h] T2. #n #h #G #L #V #T1 #T2 * /3 width=3 by cpg_eps, isrt_plus_O2, ex2_intro/ @@ -83,6 +89,7 @@ lemma cpm_ee: ∀n,h,G,L,V1,V2,T. ⦃G, L⦄ ⊢ V1 ➡[n, h] V2 → ⦃G, L⦄ /3 width=3 by cpg_ee, isrt_succ, ex2_intro/ qed. +(* Basic_2A1: includes: cpr_beta *) lemma cpm_beta: ∀n,h,p,G,L,V1,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V2 → ⦃G, L⦄ ⊢ W1 ➡[h] W2 → ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[n, h] T2 → ⦃G, L⦄ ⊢ ⓐV1.ⓛ{p}W1.T1 ➡[n, h] ⓓ{p}ⓝW2.V2.T2. @@ -90,6 +97,7 @@ lemma cpm_beta: ∀n,h,p,G,L,V1,V2,W1,W2,T1,T2. /6 width=7 by cpg_beta, isrt_plus_O2, isrt_plus, isr_shift, ex2_intro/ qed. +(* Basic_2A1: includes: cpr_theta *) lemma cpm_theta: ∀n,h,p,G,L,V1,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V → ⬆*[1] V ≡ V2 → ⦃G, L⦄ ⊢ W1 ➡[h] W2 → ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[n, h] T2 → @@ -100,10 +108,12 @@ qed. (* Basic properties on r-transition *****************************************) +(* Basic_1: includes by definition: pr0_refl *) (* Basic_2A1: includes: cpr_atom *) lemma cpr_refl: ∀h,G,L. reflexive … (cpm 0 h G L). /2 width=3 by ex2_intro/ qed. +(* Basic_1: was: pr2_head_1 *) lemma cpr_pair_sn: ∀h,I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h] V2 → ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡[h] ②{I}V2.T. #h #I #G #L #V1 #V2 * @@ -173,6 +183,7 @@ lemma cpm_inv_gref1: ∀n,h,G,L,T2,l. ⦃G, L⦄ ⊢ §l ➡[n, h] T2 → T2 = #H1 #H2 destruct /3 width=1 by isrt_inv_00, conj/ qed-. +(* Basic_2A1: includes: cpr_inv_bind1 *) lemma cpm_inv_bind1: ∀n,h,p,I,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓑ{p,I}V1.T1 ➡[n, h] U2 → ( ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V2 & ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡[n, h] T2 & U2 = ⓑ{p,I}V2.T2 @@ -189,6 +200,8 @@ lemma cpm_inv_bind1: ∀n,h,p,I,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓑ{p,I}V1.T1 ➡[n ] qed-. +(* Basic_1: includes: pr0_gen_abbr pr2_gen_abbr *) +(* Basic_2A1: includes: cpr_inv_abbr1 *) lemma cpm_inv_abbr1: ∀n,h,p,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{p}V1.T1 ➡[n, h] U2 → ( ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[n, h] T2 & U2 = ⓓ{p}V2.T2 @@ -204,6 +217,8 @@ lemma cpm_inv_abbr1: ∀n,h,p,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{p}V1.T1 ➡[n, h] ] qed-. +(* Basic_1: includes: pr0_gen_abst pr2_gen_abst *) +(* Basic_2A1: includes: cpr_inv_abst1 *) lemma cpm_inv_abst1: ∀n,h,p,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{p}V1.T1 ➡[n, h] U2 → ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ➡[n, h] T2 & U2 = ⓛ{p}V2.T2. @@ -254,6 +269,8 @@ lemma cpm_inv_flat1: ∀n,h,I,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ➡[n, h] ] qed-. +(* Basic_1: includes: pr0_gen_appl pr2_gen_appl *) +(* Basic_2A1: includes: cpr_inv_appl1 *) lemma cpm_inv_appl1: ∀n,h,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓐ V1.U1 ➡[n, h] U2 → ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V2 & ⦃G, L⦄ ⊢ U1 ➡[n, h] T2 & U2 = ⓐV2.T2 @@ -305,6 +322,7 @@ qed-. (* Basic forward lemmas *****************************************************) +(* Basic_2A1: includes: cpr_fwd_bind1_minus *) lemma cpm_fwd_bind1_minus: ∀n,h,I,G,L,V1,T1,T. ⦃G, L⦄ ⊢ -ⓑ{I}V1.T1 ➡[n, h] T → ∀p. ∃∃V2,T2. ⦃G, L⦄ ⊢ ⓑ{p,I}V1.T1 ➡[n, h] ⓑ{p,I}V2.T2 & T = -ⓑ{I}V2.T2. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpr.etc b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpr.etc new file mode 100644 index 000000000..dee79d193 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpr.etc @@ -0,0 +1,87 @@ +(* Basic_1: includes: pr2_delta1 *) +| cpr_delta: ∀G,L,K,V,V2,W2,i. + ⬇[i] L ≡ K. ⓓV → cpr G K V V2 → + ⬆[0, i + 1] V2 ≡ W2 → cpr G L (#i) W2 + +lemma cpr_cpx: ∀h,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ T1 ➡[h] T2. +#h #o #G #L #T1 #T2 #H elim H -L -T1 -T2 +/2 width=7 by cpx_delta, cpx_bind, cpx_flat, cpx_zeta, cpx_eps, cpx_beta, cpx_theta/ +qed. + +lemma lsubr_cpr_trans: ∀G. lsub_trans … (cpr G) lsubr. +#G #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 +[ // +| #G #L1 #K1 #V1 #V2 #W2 #i #HLK1 #_ #HVW2 #IHV12 #L2 #HL12 + elim (lsubr_fwd_drop2_abbr … HL12 … HLK1) -L1 * + /3 width=6 by cpr_delta/ +|3,7: /4 width=1 by lsubr_pair, cpr_bind, cpr_beta/ +|4,6: /3 width=1 by cpr_flat, cpr_eps/ +|5,8: /4 width=3 by lsubr_pair, cpr_zeta, cpr_theta/ +] +qed-. + +(* Basic_1: was by definition: pr2_free *) +lemma tpr_cpr: ∀G,T1,T2. ⦃G, ⋆⦄ ⊢ T1 ➡ T2 → ∀L. ⦃G, L⦄ ⊢ T1 ➡ T2. +#G #T1 #T2 #HT12 #L +lapply (lsubr_cpr_trans … HT12 L ?) // +qed. + +lemma cpr_delift: ∀G,K,V,T1,L,l. ⬇[l] L ≡ (K.ⓓV) → + ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ➡ T2 & ⬆[l, 1] T ≡ T2. +#G #K #V #T1 elim T1 -T1 +[ * /2 width=4 by cpr_atom, lift_sort, lift_gref, ex2_2_intro/ + #i #L #l #HLK elim (lt_or_eq_or_gt i l) + #Hil [1,3: /4 width=4 by lift_lref_ge_minus, lift_lref_lt, ylt_inj, yle_inj, ex2_2_intro/ ] + destruct + elim (lift_total V 0 (i+1)) #W #HVW + elim (lift_split … HVW i i) /3 width=6 by cpr_delta, ex2_2_intro/ +| * [ #a ] #I #W1 #U1 #IHW1 #IHU1 #L #l #HLK + elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2 + [ elim (IHU1 (L. ⓑ{I}W1) (l+1)) -IHU1 /3 width=9 by drop_drop, cpr_bind, lift_bind, ex2_2_intro/ + | elim (IHU1 … HLK) -IHU1 -HLK /3 width=8 by cpr_flat, lift_flat, ex2_2_intro/ + ] +] +qed-. + +fact lstas_cpr_aux: ∀h,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*[h, d] T2 → + d = 0 → ⦃G, L⦄ ⊢ T1 ➡ T2. +#h #G #L #T1 #T2 #d #H elim H -G -L -T1 -T2 -d +/3 width=1 by cpr_eps, cpr_flat, cpr_bind/ +[ #G #L #K #V1 #V2 #W2 #i #d #HLK #_ #HVW2 #IHV12 #H destruct + /3 width=6 by cpr_delta/ +| #G #L #K #V1 #V2 #W2 #i #d #_ #_ #_ #_ (max_O2 n) /2 width=1 by isrt_max/ +qed. + +(* Inversion properties with test for constrained rt-transition counter *****) + +lemma isrt_inv_max: ∀n,c1,c2. 𝐑𝐓⦃n, c1 ∨ c2⦄ → + ∃∃n1,n2. 𝐑𝐓⦃n1, c1⦄ & 𝐑𝐓⦃n2, c2⦄ & (n1 ∨ n2) = n. +#n #c1 #c2 * #ri #rs #H +elim (max_inv_dx … H) -H #ri1 #rs1 #ti1 #ts1 #ri2 #rs2 #ti2 #ts2 #_ #_ #H1 #H2 #H3 #H4 +elim (max_inv_O3 … H1) -H1 /3 width=5 by ex3_2_intro, ex1_2_intro/ +qed-. + +lemma isrt_inv_max_O_dx: ∀n,c1,c2. 𝐑𝐓⦃n, c1 ∨ c2⦄ → 𝐑𝐓⦃0, c2⦄ → 𝐑𝐓⦃n, c1⦄. +#n #c1 #c2 #H #H2 +elim (isrt_inv_max … H) -H #n1 #n2 #Hn1 #Hn2 #H destruct +lapply (isrt_mono … Hn2 H2) -c2 #H destruct // +qed-. + +(* Properties with shift ****************************************************) +(* +lemma max_shift: ∀c1,c2. (↓c1) ∨ (↓c2) = ↓(c1∨c2). +* #ri1 #rs1 #ti1 #ts1 * #ri2 #rs2 #ti2 #ts2 +