X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_transition%2Fcpr.ma;h=fb1f479dd4d3f02789abadcf140350596c886aeb;hb=50a9ed8c6207145fccf59e6a5dbbff935cd2c6d7;hp=fd5163073baa4891db4e5c0d4a6adca55f699335;hpb=747b42f3b9aac5487047f57742f1fcf05b56b57d;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpr.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpr.ma index fd5163073..fb1f479dd 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpr.ma @@ -12,6 +12,7 @@ (* *) (**************************************************************************) +include "ground_2/insert_eq/insert_eq_0.ma". include "basic_2/rt_transition/cpm.ma". (* CONTEXT-SENSITIVE PARALLEL R-TRANSITION FOR TERMS ************************) @@ -36,32 +37,34 @@ qed. lemma cpr_inv_atom1: ∀h,J,G,L,T2. ⦃G, L⦄ ⊢ ⓪{J} ➡[h] T2 → ∨∨ T2 = ⓪{J} - | ∃∃K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[h] V2 & ⬆*[1] V2 ≡ T2 & + | ∃∃K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[h] V2 & ⬆*[1] V2 ≘ T2 & L = K.ⓓV1 & J = LRef 0 - | ∃∃I,K,T,i. ⦃G, K⦄ ⊢ #i ➡[h] T & ⬆*[1] T ≡ T2 & - L = K.ⓘ{I} & J = LRef (⫯i). + | ∃∃I,K,T,i. ⦃G, K⦄ ⊢ #i ➡[h] T & ⬆*[1] T ≘ T2 & + L = K.ⓘ{I} & J = LRef (↑i). #h #J #G #L #T2 #H elim (cpm_inv_atom1 … H) -H * -/3 width=8 by tri_lt, or3_intro0, or3_intro1, or3_intro2, ex4_4_intro, ex4_3_intro/ -#n #_ #_ #H destruct +[2,4:|*: /3 width=8 by or3_intro0, or3_intro1, or3_intro2, ex4_4_intro, ex4_3_intro/ ] +[ #n #_ #_ #H destruct +| #n #K #V1 #V2 #_ #_ #_ #_ #H destruct +] qed-. (* Basic_1: includes: pr0_gen_sort pr2_gen_sort *) lemma cpr_inv_sort1: ∀h,G,L,T2,s. ⦃G, L⦄ ⊢ ⋆s ➡[h] T2 → T2 = ⋆s. -#h #G #L #T2 #s #H elim (cpm_inv_sort1 … H) -H * // #_ #H destruct +#h #G #L #T2 #s #H elim (cpm_inv_sort1 … H) -H // qed-. lemma cpr_inv_zero1: ∀h,G,L,T2. ⦃G, L⦄ ⊢ #0 ➡[h] T2 → ∨∨ T2 = #0 - | ∃∃K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[h] V2 & ⬆*[1] V2 ≡ T2 & + | ∃∃K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[h] V2 & ⬆*[1] V2 ≘ T2 & L = K.ⓓV1. #h #G #L #T2 #H elim (cpm_inv_zero1 … H) -H * /3 width=6 by ex3_3_intro, or_introl, or_intror/ #n #K #V1 #V2 #_ #_ #_ #H destruct qed-. -lemma cpr_inv_lref1: ∀h,G,L,T2,i. ⦃G, L⦄ ⊢ #⫯i ➡[h] T2 → - ∨∨ T2 = #(⫯i) - | ∃∃I,K,T. ⦃G, K⦄ ⊢ #i ➡[h] T & ⬆*[1] T ≡ T2 & L = K.ⓘ{I}. +lemma cpr_inv_lref1: ∀h,G,L,T2,i. ⦃G, L⦄ ⊢ #↑i ➡[h] T2 → + ∨∨ T2 = #(↑i) + | ∃∃I,K,T. ⦃G, K⦄ ⊢ #i ➡[h] T & ⬆*[1] T ≘ T2 & L = K.ⓘ{I}. #h #G #L #T2 #i #H elim (cpm_inv_lref1 … H) -H * /3 width=6 by ex3_3_intro, or_introl, or_intror/ qed-. @@ -86,7 +89,7 @@ lemma cpr_inv_flat1: ∀h,I,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ➡[h] U2 | ∃∃p,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V2 & ⦃G, L⦄ ⊢ W1 ➡[h] W2 & ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h] T2 & U1 = ⓛ{p}W1.T1 & U2 = ⓓ{p}ⓝW2.V2.T2 & I = Appl - | ∃∃p,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V & ⬆*[1] V ≡ V2 & + | ∃∃p,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V & ⬆*[1] V ≘ V2 & ⦃G, L⦄ ⊢ W1 ➡[h] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h] T2 & U1 = ⓓ{p}W1.T1 & U2 = ⓓ{p}W2.ⓐV2.T2 & I = Appl. @@ -98,6 +101,39 @@ lemma cpr_inv_flat1: ∀h,I,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ➡[h] U2 ] qed-. +(* Basic eliminators ********************************************************) + +lemma cpr_ind (h): ∀Q:relation4 genv lenv term term. + (∀I,G,L. Q G L (⓪{I}) (⓪{I})) → + (∀G,K,V1,V2,W2. ⦃G, K⦄ ⊢ V1 ➡[h] V2 → Q G K V1 V2 → + ⬆*[1] V2 ≘ W2 → Q G (K.ⓓV1) (#0) W2 + ) → (∀I,G,K,T,U,i. ⦃G, K⦄ ⊢ #i ➡[h] T → Q G K (#i) T → + ⬆*[1] T ≘ U → Q G (K.ⓘ{I}) (#↑i) (U) + ) → (∀p,I,G,L,V1,V2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V2 → ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡[h] T2 → + Q G L V1 V2 → Q G (L.ⓑ{I}V1) T1 T2 → Q G L (ⓑ{p,I}V1.T1) (ⓑ{p,I}V2.T2) + ) → (∀I,G,L,V1,V2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V2 → ⦃G, L⦄ ⊢ T1 ➡[h] T2 → + Q G L V1 V2 → Q G L T1 T2 → Q G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) + ) → (∀G,L,V,T1,T,T2. ⬆*[1] T ≘ T1 → ⦃G, L⦄ ⊢ T ➡[h] T2 → + Q G L T T2 → Q G L (+ⓓV.T1) T2 + ) → (∀G,L,V,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h] T2 → Q G L T1 T2 → + Q G L (ⓝV.T1) T2 + ) → (∀p,G,L,V1,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V2 → ⦃G, L⦄ ⊢ W1 ➡[h] W2 → ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h] T2 → + Q G L V1 V2 → Q G L W1 W2 → Q G (L.ⓛW1) T1 T2 → + Q G L (ⓐV1.ⓛ{p}W1.T1) (ⓓ{p}ⓝW2.V2.T2) + ) → (∀p,G,L,V1,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V → ⦃G, L⦄ ⊢ W1 ➡[h] W2 → ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h] T2 → + Q G L V1 V → Q G L W1 W2 → Q G (L.ⓓW1) T1 T2 → + ⬆*[1] V ≘ V2 → Q G L (ⓐV1.ⓓ{p}W1.T1) (ⓓ{p}W2.ⓐV2.T2) + ) → + ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h] T2 → Q G L T1 T2. +#h #Q #IH1 #IH2 #IH3 #IH4 #IH5 #IH6 #IH7 #IH8 #IH9 #G #L #T1 #T2 +@(insert_eq_0 … 0) #n #H +@(cpm_ind … H) -G -L -T1 -T2 -n [2,4,11:|*: /3 width=4 by/ ] +[ #G #L #s #H destruct +| #n #G #K #V1 #V2 #W2 #_ #_ #_ #H destruct +| #n #G #L #U1 #U2 #T #_ #_ #H destruct +] +qed-. + (* Basic_1: removed theorems 12: pr0_subst0_back pr0_subst0_fwd pr0_subst0 pr0_delta1