X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Freduction%2Fcpx.ma;h=f7fd702e50cbaf7852b0d2356bb886d1b933c22b;hb=93bba1c94779e83184d111cd077d4167e42a74aa;hp=840bf0866c06eb991a1f43207a72fb7cb299ca7b;hpb=9a023f554e56d6edbbb2eeaf17ce61e31857ef4a;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx.ma index 840bf0866..f7fd702e5 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx.ma @@ -19,39 +19,39 @@ include "basic_2/reduction/cpr.ma". (* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************) (* avtivate genv *) -inductive cpx (h) (g): relation4 genv lenv term term ≝ -| cpx_atom : ∀I,G,L. cpx h g G L (⓪{I}) (⓪{I}) -| cpx_st : ∀G,L,k,d. deg h g k (d+1) → cpx h g G L (⋆k) (⋆(next h k)) +inductive cpx (h) (o): relation4 genv lenv term term ≝ +| cpx_atom : ∀I,G,L. cpx h o G L (⓪{I}) (⓪{I}) +| cpx_st : ∀G,L,s,d. deg h o s (d+1) → cpx h o G L (⋆s) (⋆(next h s)) | cpx_delta: ∀I,G,L,K,V,V2,W2,i. - ⬇[i] L ≡ K.ⓑ{I}V → cpx h g G K V V2 → - ⬆[0, i+1] V2 ≡ W2 → cpx h g G L (#i) W2 + ⬇[i] L ≡ K.ⓑ{I}V → cpx h o G K V V2 → + ⬆[0, i+1] V2 ≡ W2 → cpx h o G L (#i) W2 | cpx_bind : ∀a,I,G,L,V1,V2,T1,T2. - cpx h g G L V1 V2 → cpx h g G (L.ⓑ{I}V1) T1 T2 → - cpx h g G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2) + cpx h o G L V1 V2 → cpx h o G (L.ⓑ{I}V1) T1 T2 → + cpx h o G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2) | cpx_flat : ∀I,G,L,V1,V2,T1,T2. - cpx h g G L V1 V2 → cpx h g G L T1 T2 → - cpx h g G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) -| cpx_zeta : ∀G,L,V,T1,T,T2. cpx h g G (L.ⓓV) T1 T → - ⬆[0, 1] T2 ≡ T → cpx h g G L (+ⓓV.T1) T2 -| cpx_eps : ∀G,L,V,T1,T2. cpx h g G L T1 T2 → cpx h g G L (ⓝV.T1) T2 -| cpx_ct : ∀G,L,V1,V2,T. cpx h g G L V1 V2 → cpx h g G L (ⓝV1.T) V2 + cpx h o G L V1 V2 → cpx h o G L T1 T2 → + cpx h o G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) +| cpx_zeta : ∀G,L,V,T1,T,T2. cpx h o G (L.ⓓV) T1 T → + ⬆[0, 1] T2 ≡ T → cpx h o G L (+ⓓV.T1) T2 +| cpx_eps : ∀G,L,V,T1,T2. cpx h o G L T1 T2 → cpx h o G L (ⓝV.T1) T2 +| cpx_ct : ∀G,L,V1,V2,T. cpx h o G L V1 V2 → cpx h o G L (ⓝV1.T) V2 | cpx_beta : ∀a,G,L,V1,V2,W1,W2,T1,T2. - cpx h g G L V1 V2 → cpx h g G L W1 W2 → cpx h g G (L.ⓛW1) T1 T2 → - cpx h g G L (ⓐV1.ⓛ{a}W1.T1) (ⓓ{a}ⓝW2.V2.T2) + cpx h o G L V1 V2 → cpx h o G L W1 W2 → cpx h o G (L.ⓛW1) T1 T2 → + cpx h o G L (ⓐV1.ⓛ{a}W1.T1) (ⓓ{a}ⓝW2.V2.T2) | cpx_theta: ∀a,G,L,V1,V,V2,W1,W2,T1,T2. - cpx h g G L V1 V → ⬆[0, 1] V ≡ V2 → cpx h g G L W1 W2 → - cpx h g G (L.ⓓW1) T1 T2 → - cpx h g G L (ⓐV1.ⓓ{a}W1.T1) (ⓓ{a}W2.ⓐV2.T2) + cpx h o G L V1 V → ⬆[0, 1] V ≡ V2 → cpx h o G L W1 W2 → + cpx h o G (L.ⓓW1) T1 T2 → + cpx h o G L (ⓐV1.ⓓ{a}W1.T1) (ⓓ{a}W2.ⓐV2.T2) . interpretation "context-sensitive extended parallel reduction (term)" - 'PRed h g G L T1 T2 = (cpx h g G L T1 T2). + 'PRed h o G L T1 T2 = (cpx h o G L T1 T2). (* Basic properties *********************************************************) -lemma lsubr_cpx_trans: ∀h,g,G. lsub_trans … (cpx h g G) lsubr. -#h #g #G #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 +lemma lsubr_cpx_trans: ∀h,o,G. lsub_trans … (cpx h o G) lsubr. +#h #o #G #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 [ // | /2 width=2 by cpx_st/ | #I #G #L1 #K1 #V1 #V2 #W2 #i #HLK1 #_ #HVW2 #IHV12 #L2 #HL12 @@ -64,23 +64,23 @@ lemma lsubr_cpx_trans: ∀h,g,G. lsub_trans … (cpx h g G) lsubr. qed-. (* Note: this is "∀h,g,L. reflexive … (cpx h g L)" *) -lemma cpx_refl: ∀h,g,G,T,L. ⦃G, L⦄ ⊢ T ➡[h, g] T. -#h #g #G #T elim T -T // * /2 width=1 by cpx_bind, cpx_flat/ +lemma cpx_refl: ∀h,o,G,T,L. ⦃G, L⦄ ⊢ T ➡[h, o] T. +#h #o #G #T elim T -T // * /2 width=1 by cpx_bind, cpx_flat/ qed. -lemma cpr_cpx: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2. -#h #g #G #L #T1 #T2 #H elim H -L -T1 -T2 +lemma cpr_cpx: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ T1 ➡[h, o] 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 cpx_pair_sn: ∀h,g,I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 → - ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡[h, g] ②{I}V2.T. -#h #g * /2 width=1 by cpx_bind, cpx_flat/ +lemma cpx_pair_sn: ∀h,o,I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 → + ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡[h, o] ②{I}V2.T. +#h #o * /2 width=1 by cpx_bind, cpx_flat/ qed. -lemma cpx_delift: ∀h,g,I,G,K,V,T1,L,l. ⬇[l] L ≡ (K.ⓑ{I}V) → - ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 & ⬆[l, 1] T ≡ T2. -#h #g #I #G #K #V #T1 elim T1 -T1 +lemma cpx_delift: ∀h,o,I,G,K,V,T1,L,l. ⬇[l] L ≡ (K.ⓑ{I}V) → + ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 & ⬆[l, 1] T ≡ T2. +#h #o #I #G #K #V #T1 elim T1 -T1 [ * #i #L #l /2 width=4 by cpx_atom, lift_sort, lift_gref, ex2_2_intro/ elim (lt_or_eq_or_gt i l) #Hil [1,3: /4 width=4 by cpx_atom, lift_lref_ge_minus, lift_lref_lt, ylt_inj, yle_inj, ex2_2_intro/ ] destruct @@ -96,14 +96,14 @@ qed-. (* Basic inversion lemmas ***************************************************) -fact cpx_inv_atom1_aux: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → ∀J. T1 = ⓪{J} → +fact cpx_inv_atom1_aux: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → ∀J. T1 = ⓪{J} → ∨∨ T2 = ⓪{J} - | ∃∃k,d. deg h g k (d+1) & T2 = ⋆(next h k) & J = Sort k - | ∃∃I,K,V,V2,i. ⬇[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, g] V2 & + | ∃∃s,d. deg h o s (d+1) & T2 = ⋆(next h s) & J = Sort s + | ∃∃I,K,V,V2,i. ⬇[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, o] V2 & ⬆[O, i+1] V2 ≡ T2 & J = LRef i. -#G #h #g #L #T1 #T2 * -L -T1 -T2 +#G #h #o #L #T1 #T2 * -L -T1 -T2 [ #I #G #L #J #H destruct /2 width=1 by or3_intro0/ -| #G #L #k #d #Hkd #J #H destruct /3 width=5 by or3_intro1, ex3_2_intro/ +| #G #L #s #d #Hkd #J #H destruct /3 width=5 by or3_intro1, ex3_2_intro/ | #I #G #L #K #V #V2 #T2 #i #HLK #HV2 #HVT2 #J #H destruct /3 width=9 by or3_intro2, ex4_5_intro/ | #a #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #H destruct | #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #H destruct @@ -115,57 +115,57 @@ fact cpx_inv_atom1_aux: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → ∀ ] qed-. -lemma cpx_inv_atom1: ∀h,g,J,G,L,T2. ⦃G, L⦄ ⊢ ⓪{J} ➡[h, g] T2 → +lemma cpx_inv_atom1: ∀h,o,J,G,L,T2. ⦃G, L⦄ ⊢ ⓪{J} ➡[h, o] T2 → ∨∨ T2 = ⓪{J} - | ∃∃k,d. deg h g k (d+1) & T2 = ⋆(next h k) & J = Sort k - | ∃∃I,K,V,V2,i. ⬇[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, g] V2 & + | ∃∃s,d. deg h o s (d+1) & T2 = ⋆(next h s) & J = Sort s + | ∃∃I,K,V,V2,i. ⬇[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, o] V2 & ⬆[O, i+1] V2 ≡ T2 & J = LRef i. /2 width=3 by cpx_inv_atom1_aux/ qed-. -lemma cpx_inv_sort1: ∀h,g,G,L,T2,k. ⦃G, L⦄ ⊢ ⋆k ➡[h, g] T2 → T2 = ⋆k ∨ - ∃∃d. deg h g k (d+1) & T2 = ⋆(next h k). -#h #g #G #L #T2 #k #H +lemma cpx_inv_sort1: ∀h,o,G,L,T2,s. ⦃G, L⦄ ⊢ ⋆s ➡[h, o] T2 → T2 = ⋆s ∨ + ∃∃d. deg h o s (d+1) & T2 = ⋆(next h s). +#h #o #G #L #T2 #s #H elim (cpx_inv_atom1 … H) -H /2 width=1 by or_introl/ * -[ #k0 #d0 #Hkd0 #H1 #H2 destruct /3 width=4 by ex2_intro, or_intror/ +[ #s0 #d0 #Hkd0 #H1 #H2 destruct /3 width=4 by ex2_intro, or_intror/ | #I #K #V #V2 #i #_ #_ #_ #H destruct ] qed-. -lemma cpx_inv_lref1: ∀h,g,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, g] T2 → +lemma cpx_inv_lref1: ∀h,o,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, o] T2 → T2 = #i ∨ - ∃∃I,K,V,V2. ⬇[i] L ≡ K. ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, g] V2 & + ∃∃I,K,V,V2. ⬇[i] L ≡ K. ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, o] V2 & ⬆[O, i+1] V2 ≡ T2. -#h #g #G #L #T2 #i #H +#h #o #G #L #T2 #i #H elim (cpx_inv_atom1 … H) -H /2 width=1 by or_introl/ * -[ #k #d #_ #_ #H destruct +[ #s #d #_ #_ #H destruct | #I #K #V #V2 #j #HLK #HV2 #HVT2 #H destruct /3 width=7 by ex3_4_intro, or_intror/ ] qed-. -lemma cpx_inv_lref1_ge: ∀h,g,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, g] T2 → |L| ≤ i → T2 = #i. -#h #g #G #L #T2 #i #H elim (cpx_inv_lref1 … H) -H // * +lemma cpx_inv_lref1_ge: ∀h,o,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, o] T2 → |L| ≤ i → T2 = #i. +#h #o #G #L #T2 #i #H elim (cpx_inv_lref1 … H) -H // * #I #K #V1 #V2 #HLK #_ #_ #HL -h -G -V2 lapply (drop_fwd_length_lt2 … HLK) -K -I -V1 #H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/ qed-. -lemma cpx_inv_gref1: ∀h,g,G,L,T2,p. ⦃G, L⦄ ⊢ §p ➡[h, g] T2 → T2 = §p. -#h #g #G #L #T2 #p #H +lemma cpx_inv_gref1: ∀h,o,G,L,T2,p. ⦃G, L⦄ ⊢ §p ➡[h, o] T2 → T2 = §p. +#h #o #G #L #T2 #p #H elim (cpx_inv_atom1 … H) -H // * -[ #k #d #_ #_ #H destruct +[ #s #d #_ #_ #H destruct | #I #K #V #V2 #i #_ #_ #_ #H destruct ] qed-. -fact cpx_inv_bind1_aux: ∀h,g,G,L,U1,U2. ⦃G, L⦄ ⊢ U1 ➡[h, g] U2 → +fact cpx_inv_bind1_aux: ∀h,o,G,L,U1,U2. ⦃G, L⦄ ⊢ U1 ➡[h, o] U2 → ∀a,J,V1,T1. U1 = ⓑ{a,J}V1.T1 → ( - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L.ⓑ{J}V1⦄ ⊢ T1 ➡[h, g] T2 & + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L.ⓑ{J}V1⦄ ⊢ T1 ➡[h, o] T2 & U2 = ⓑ{a,J}V2.T2 ) ∨ - ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, g] T & ⬆[0, 1] U2 ≡ T & + ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, o] T & ⬆[0, 1] U2 ≡ T & a = true & J = Abbr. -#h #g #G #L #U1 #U2 * -L -U1 -U2 +#h #o #G #L #U1 #U2 * -L -U1 -U2 [ #I #G #L #b #J #W #U1 #H destruct -| #G #L #k #d #_ #b #J #W #U1 #H destruct +| #G #L #s #d #_ #b #J #W #U1 #H destruct | #I #G #L #K #V #V2 #W2 #i #_ #_ #_ #b #J #W #U1 #H destruct | #a #I #G #L #V1 #V2 #T1 #T2 #HV12 #HT12 #b #J #W #U1 #H destruct /3 width=5 by ex3_2_intro, or_introl/ | #I #G #L #V1 #V2 #T1 #T2 #_ #_ #b #J #W #U1 #H destruct @@ -177,50 +177,50 @@ fact cpx_inv_bind1_aux: ∀h,g,G,L,U1,U2. ⦃G, L⦄ ⊢ U1 ➡[h, g] U2 → ] qed-. -lemma cpx_inv_bind1: ∀h,g,a,I,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡[h, g] U2 → ( - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡[h, g] T2 & +lemma cpx_inv_bind1: ∀h,o,a,I,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡[h, o] U2 → ( + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡[h, o] T2 & U2 = ⓑ{a,I} V2. T2 ) ∨ - ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, g] T & ⬆[0, 1] U2 ≡ T & + ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, o] T & ⬆[0, 1] U2 ≡ T & a = true & I = Abbr. /2 width=3 by cpx_inv_bind1_aux/ qed-. -lemma cpx_inv_abbr1: ∀h,g,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡[h, g] U2 → ( - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, g] T2 & +lemma cpx_inv_abbr1: ∀h,o,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡[h, o] U2 → ( + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, o] T2 & U2 = ⓓ{a} V2. T2 ) ∨ - ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, g] T & ⬆[0, 1] U2 ≡ T & a = true. -#h #g #a #G #L #V1 #T1 #U2 #H + ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, o] T & ⬆[0, 1] U2 ≡ T & a = true. +#h #o #a #G #L #V1 #T1 #U2 #H elim (cpx_inv_bind1 … H) -H * /3 width=5 by ex3_2_intro, ex3_intro, or_introl, or_intror/ qed-. -lemma cpx_inv_abst1: ∀h,g,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 ➡[h, g] U2 → - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ➡[h, g] T2 & +lemma cpx_inv_abst1: ∀h,o,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 ➡[h, o] U2 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ➡[h, o] T2 & U2 = ⓛ{a} V2. T2. -#h #g #a #G #L #V1 #T1 #U2 #H +#h #o #a #G #L #V1 #T1 #U2 #H elim (cpx_inv_bind1 … H) -H * [ /3 width=5 by ex3_2_intro/ | #T #_ #_ #_ #H destruct ] qed-. -fact cpx_inv_flat1_aux: ∀h,g,G,L,U,U2. ⦃G, L⦄ ⊢ U ➡[h, g] U2 → +fact cpx_inv_flat1_aux: ∀h,o,G,L,U,U2. ⦃G, L⦄ ⊢ U ➡[h, o] U2 → ∀J,V1,U1. U = ⓕ{J}V1.U1 → - ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, g] T2 & + ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, o] T2 & U2 = ⓕ{J}V2.T2 - | (⦃G, L⦄ ⊢ U1 ➡[h, g] U2 ∧ J = Cast) - | (⦃G, L⦄ ⊢ V1 ➡[h, g] U2 ∧ J = Cast) - | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, g] W2 & - ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, g] T2 & + | (⦃G, L⦄ ⊢ U1 ➡[h, o] U2 ∧ J = Cast) + | (⦃G, L⦄ ⊢ V1 ➡[h, o] U2 ∧ J = Cast) + | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 & + ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, o] T2 & U1 = ⓛ{a}W1.T1 & U2 = ⓓ{a}ⓝW2.V2.T2 & J = Appl - | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V & ⬆[0,1] V ≡ V2 & - ⦃G, L⦄ ⊢ W1 ➡[h, g] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, g] T2 & + | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V & ⬆[0,1] V ≡ V2 & + ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, o] T2 & U1 = ⓓ{a}W1.T1 & U2 = ⓓ{a}W2.ⓐV2.T2 & J = Appl. -#h #g #G #L #U #U2 * -L -U -U2 +#h #o #G #L #U #U2 * -L -U -U2 [ #I #G #L #J #W #U1 #H destruct -| #G #L #k #d #_ #J #W #U1 #H destruct +| #G #L #s #d #_ #J #W #U1 #H destruct | #I #G #L #K #V #V2 #W2 #i #_ #_ #_ #J #W #U1 #H destruct | #a #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #W #U1 #H destruct | #I #G #L #V1 #V2 #T1 #T2 #HV12 #HT12 #J #W #U1 #H destruct /3 width=5 by or5_intro0, ex3_2_intro/ @@ -232,31 +232,31 @@ fact cpx_inv_flat1_aux: ∀h,g,G,L,U,U2. ⦃G, L⦄ ⊢ U ➡[h, g] U2 → ] qed-. -lemma cpx_inv_flat1: ∀h,g,I,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ➡[h, g] U2 → - ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, g] T2 & +lemma cpx_inv_flat1: ∀h,o,I,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ➡[h, o] U2 → + ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, o] T2 & U2 = ⓕ{I} V2. T2 - | (⦃G, L⦄ ⊢ U1 ➡[h, g] U2 ∧ I = Cast) - | (⦃G, L⦄ ⊢ V1 ➡[h, g] U2 ∧ I = Cast) - | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, g] W2 & - ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, g] T2 & + | (⦃G, L⦄ ⊢ U1 ➡[h, o] U2 ∧ I = Cast) + | (⦃G, L⦄ ⊢ V1 ➡[h, o] U2 ∧ I = Cast) + | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 & + ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, o] T2 & U1 = ⓛ{a}W1.T1 & U2 = ⓓ{a}ⓝW2.V2.T2 & I = Appl - | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V & ⬆[0,1] V ≡ V2 & - ⦃G, L⦄ ⊢ W1 ➡[h, g] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, g] T2 & + | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V & ⬆[0,1] V ≡ V2 & + ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, o] T2 & U1 = ⓓ{a}W1.T1 & U2 = ⓓ{a}W2.ⓐV2.T2 & I = Appl. /2 width=3 by cpx_inv_flat1_aux/ qed-. -lemma cpx_inv_appl1: ∀h,g,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓐ V1.U1 ➡[h, g] U2 → - ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, g] T2 & +lemma cpx_inv_appl1: ∀h,o,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓐ V1.U1 ➡[h, o] U2 → + ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, o] T2 & U2 = ⓐ V2. T2 - | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, g] W2 & - ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, g] T2 & + | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 & + ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, o] T2 & U1 = ⓛ{a}W1.T1 & U2 = ⓓ{a}ⓝW2.V2.T2 - | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V & ⬆[0,1] V ≡ V2 & - ⦃G, L⦄ ⊢ W1 ➡[h, g] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, g] T2 & + | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V & ⬆[0,1] V ≡ V2 & + ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, o] T2 & U1 = ⓓ{a}W1.T1 & U2 = ⓓ{a}W2. ⓐV2. T2. -#h #g #G #L #V1 #U1 #U2 #H elim (cpx_inv_flat1 … H) -H * +#h #o #G #L #V1 #U1 #U2 #H elim (cpx_inv_flat1 … H) -H * [ /3 width=5 by or3_intro0, ex3_2_intro/ |2,3: #_ #H destruct | /3 width=11 by or3_intro1, ex5_6_intro/ @@ -265,10 +265,10 @@ lemma cpx_inv_appl1: ∀h,g,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓐ V1.U1 ➡[h, g] U2 qed-. (* Note: the main property of simple terms *) -lemma cpx_inv_appl1_simple: ∀h,g,G,L,V1,T1,U. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡[h, g] U → 𝐒⦃T1⦄ → - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 & +lemma cpx_inv_appl1_simple: ∀h,o,G,L,V1,T1,U. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡[h, o] U → 𝐒⦃T1⦄ → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 & U = ⓐV2.T2. -#h #g #G #L #V1 #T1 #U #H #HT1 +#h #o #G #L #V1 #T1 #U #H #HT1 elim (cpx_inv_appl1 … H) -H * [ /2 width=5 by ex3_2_intro/ | #a #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #H #_ destruct @@ -278,12 +278,12 @@ elim (cpx_inv_appl1 … H) -H * ] qed-. -lemma cpx_inv_cast1: ∀h,g,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓝV1.U1 ➡[h, g] U2 → - ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, g] T2 & +lemma cpx_inv_cast1: ∀h,o,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓝV1.U1 ➡[h, o] U2 → + ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, o] T2 & U2 = ⓝ V2. T2 - | ⦃G, L⦄ ⊢ U1 ➡[h, g] U2 - | ⦃G, L⦄ ⊢ V1 ➡[h, g] U2. -#h #g #G #L #V1 #U1 #U2 #H elim (cpx_inv_flat1 … H) -H * + | ⦃G, L⦄ ⊢ U1 ➡[h, o] U2 + | ⦃G, L⦄ ⊢ V1 ➡[h, o] U2. +#h #o #G #L #V1 #U1 #U2 #H elim (cpx_inv_flat1 … H) -H * [ /3 width=5 by or3_intro0, ex3_2_intro/ |2,3: /2 width=1 by or3_intro1, or3_intro2/ | #a #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #H destruct @@ -293,10 +293,10 @@ qed-. (* Basic forward lemmas *****************************************************) -lemma cpx_fwd_bind1_minus: ∀h,g,I,G,L,V1,T1,T. ⦃G, L⦄ ⊢ -ⓑ{I}V1.T1 ➡[h, g] T → ∀b. - ∃∃V2,T2. ⦃G, L⦄ ⊢ ⓑ{b,I}V1.T1 ➡[h, g] ⓑ{b,I}V2.T2 & +lemma cpx_fwd_bind1_minus: ∀h,o,I,G,L,V1,T1,T. ⦃G, L⦄ ⊢ -ⓑ{I}V1.T1 ➡[h, o] T → ∀b. + ∃∃V2,T2. ⦃G, L⦄ ⊢ ⓑ{b,I}V1.T1 ➡[h, o] ⓑ{b,I}V2.T2 & T = -ⓑ{I}V2.T2. -#h #g #I #G #L #V1 #T1 #T #H #b +#h #o #I #G #L #V1 #T1 #T #H #b elim (cpx_inv_bind1 … H) -H * [ #V2 #T2 #HV12 #HT12 #H destruct /3 width=4 by cpx_bind, ex2_2_intro/ | #T2 #_ #_ #H destruct