X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_transition%2Fcpr.ma;h=b51753a79d40d612a97e73add01e6f3d99c81ecd;hb=d8d00d6f6694155be5be486a8239f5953efe28b7;hp=6bcfa173e8ec96dfc351c475a5820c71c14026ed;hpb=984856dbab870ddc3156040df69b1f1846cc3aaf;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 6bcfa173e..b51753a79 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,10 @@ (* *) (**************************************************************************) +include "ground_2/xoa/ex_6_6.ma". +include "ground_2/xoa/ex_7_7.ma". +include "ground_2/xoa/or_4.ma". +include "ground_2/insert_eq/insert_eq_0.ma". include "basic_2/rt_transition/cpm.ma". (* CONTEXT-SENSITIVE PARALLEL R-TRANSITION FOR TERMS ************************) @@ -21,75 +25,75 @@ include "basic_2/rt_transition/cpm.ma". (* Note: cpr_flat: does not hold in basic_1 *) (* Basic_1: includes: pr2_thin_dx *) lemma cpr_flat: ∀h,I,G,L,V1,V2,T1,T2. - ⦃G, L⦄ ⊢ V1 ➡[h] V2 → ⦃G, L⦄ ⊢ T1 ➡[h] T2 → - ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ➡[h] ⓕ{I}V2.T2. + ⦃G,L⦄ ⊢ V1 ➡[h] V2 → ⦃G,L⦄ ⊢ T1 ➡[h] T2 → + ⦃G,L⦄ ⊢ ⓕ{I}V1.T1 ➡[h] ⓕ{I}V2.T2. #h * /2 width=1 by cpm_cast, cpm_appl/ -qed. +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. +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 * /2 width=1 by cpm_bind, cpr_flat/ qed. (* Basic inversion properties ***********************************************) -lemma cpr_inv_atom1: ∀h,J,G,L,T2. ⦃G, L⦄ ⊢ ⓪{J} ➡[h] T2 → +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,V,T,i. ⦃G, K⦄ ⊢ #i ➡[h] T & ⬆*[1] T ≡ T2 & - L = K.ⓑ{I}V & 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=9 by or3_intro0, or3_intro1, or3_intro2, ex4_5_intro, ex4_3_intro/ +[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 +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 // 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 & - L = K.ⓓV1. +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 & + 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,V,T. ⦃G, K⦄ ⊢ #i ➡[h] T & ⬆*[1] T ≡ T2 & L = K.ⓑ{I}V. +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=7 by ex3_4_intro, or_introl, or_intror/ +/3 width=6 by ex3_3_intro, or_introl, or_intror/ qed-. -lemma cpr_inv_gref1: ∀h,G,L,T2,l. ⦃G, L⦄ ⊢ §l ➡[h] T2 → T2 = §l. +lemma cpr_inv_gref1: ∀h,G,L,T2,l. ⦃G,L⦄ ⊢ §l ➡[h] T2 → T2 = §l. #h #G #L #T2 #l #H elim (cpm_inv_gref1 … H) -H // qed-. (* Basic_1: includes: pr0_gen_cast pr2_gen_cast *) -lemma cpr_inv_cast1: ∀h,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓝ V1. U1 ➡[h] U2 → ( - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V2 & ⦃G, L⦄ ⊢ U1 ➡[h] T2 & - U2 = ⓝV2.T2 - ) ∨ ⦃G, L⦄ ⊢ U1 ➡[h] U2. +lemma cpr_inv_cast1: ∀h,G,L,V1,U1,U2. ⦃G,L⦄ ⊢ ⓝ V1.U1 ➡[h] U2 → + ∨∨ ∃∃V2,T2. ⦃G,L⦄ ⊢ V1 ➡[h] V2 & ⦃G,L⦄ ⊢ U1 ➡[h] T2 & + U2 = ⓝV2.T2 + | ⦃G,L⦄ ⊢ U1 ➡[h] U2. #h #G #L #V1 #U1 #U2 #H elim (cpm_inv_cast1 … H) -H /2 width=1 by or_introl, or_intror/ * #n #_ #H destruct qed-. -lemma cpr_inv_flat1: ∀h,I,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ➡[h] U2 → - ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h] V2 & ⦃G, L⦄ ⊢ U1 ➡[h] T2 & +lemma cpr_inv_flat1: ∀h,I,G,L,V1,U1,U2. ⦃G,L⦄ ⊢ ⓕ{I}V1.U1 ➡[h] U2 → + ∨∨ ∃∃V2,T2. ⦃G,L⦄ ⊢ V1 ➡[h] V2 & ⦃G,L⦄ ⊢ U1 ➡[h] T2 & U2 = ⓕ{I}V2.T2 - | (⦃G, L⦄ ⊢ U1 ➡[h] U2 ∧ I = Cast) - | ∃∃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 & + | (⦃G,L⦄ ⊢ U1 ➡[h] U2 ∧ I = Cast) + | ∃∃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 & - ⦃G, L⦄ ⊢ W1 ➡[h] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h] T2 & + | ∃∃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. #h * #G #L #V1 #U1 #U2 #H @@ -100,6 +104,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