X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Freduction%2Fcix.ma;h=fc1de0d4cf05fef518c0705ec1327ae365a04ca1;hb=f725a35c9014595293cfe43081ef11b059d5e3a7;hp=f89a5038d24294b08e94a24d43f39846c5b3158a;hpb=29973426e0227ee48368d1c24dc0c17bf2baef77;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cix.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cix.ma index f89a5038d..fc1de0d4c 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cix.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cix.ma @@ -12,50 +12,50 @@ (* *) (**************************************************************************) -include "basic_2/notation/relations/notreducible_4.ma". +include "basic_2/notation/relations/notreducible_5.ma". include "basic_2/reduction/cir.ma". include "basic_2/reduction/crx.ma". (* CONTEXT-SENSITIVE EXTENDED IRREDUCIBLE TERMS *****************************) -definition cix: ∀h. sd h → lenv → predicate term ≝ λh,g,L,T. ⦃G, L⦄ ⊢ 𝐑[h, g]⦃T⦄ → ⊥. +definition cix: ∀h. sd h → relation3 genv lenv term ≝ + λh,g,G,L,T. ⦃G, L⦄ ⊢ 𝐑[h, g]⦃T⦄ → ⊥. interpretation "context-sensitive extended irreducibility (term)" - 'NotReducible h g L T = (cix h g L T). + 'NotReducible h g G L T = (cix h g G L T). (* Basic inversion lemmas ***************************************************) -lemma cix_inv_sort: ∀h,g,L,k,l. deg h g k (l+1) → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃⋆k⦄ → ⊥. +lemma cix_inv_sort: ∀h,g,G,L,k,l. deg h g k (l+1) → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃⋆k⦄ → ⊥. /3 width=2/ qed-. -lemma cix_inv_delta: ∀h,g,I,L,K,V,i. ⇩[0, i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃#i⦄ → ⊥. +lemma cix_inv_delta: ∀h,g,I,G,L,K,V,i. ⇩[0, i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃#i⦄ → ⊥. /3 width=4/ qed-. -lemma cix_inv_ri2: ∀h,g,I,L,V,T. ri2 I → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃②{I}V.T⦄ → ⊥. +lemma cix_inv_ri2: ∀h,g,I,G,L,V,T. ri2 I → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃②{I}V.T⦄ → ⊥. /3 width=1/ qed-. -lemma cix_inv_ib2: ∀h,g,a,I,L,V,T. ib2 a I → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃ⓑ{a,I}V.T⦄ → - ⦃G, L⦄ ⊢ 𝐈[h, g]⦃V⦄ ∧ ⦃h, L.ⓑ{I}V⦄ ⊢ 𝐈[h, g]⦃T⦄. +lemma cix_inv_ib2: ∀h,g,a,I,G,L,V,T. ib2 a I → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃ⓑ{a,I}V.T⦄ → + ⦃G, L⦄ ⊢ 𝐈[h, g]⦃V⦄ ∧ ⦃G, L.ⓑ{I}V⦄ ⊢ 𝐈[h, g]⦃T⦄. /4 width=1/ qed-. -lemma cix_inv_bind: ∀h,g,a,I,L,V,T. ⦃G, L⦄ ⊢ 𝐈[h, g]⦃ⓑ{a,I}V.T⦄ → - ∧∧ ⦃G, L⦄ ⊢ 𝐈[h, g]⦃V⦄ & ⦃h, L.ⓑ{I}V⦄ ⊢ 𝐈[h, g]⦃T⦄ & ib2 a I. +lemma cix_inv_bind: ∀h,g,a,I,G,L,V,T. ⦃G, L⦄ ⊢ 𝐈[h, g]⦃ⓑ{a,I}V.T⦄ → + ∧∧ ⦃G, L⦄ ⊢ 𝐈[h, g]⦃V⦄ & ⦃G, L.ⓑ{I}V⦄ ⊢ 𝐈[h, g]⦃T⦄ & ib2 a I. #h #g #a * [ elim a -a ] -[ #L #V #T #H elim H -H /3 width=1/ -|*: #L #V #T #H elim (cix_inv_ib2 … H) -H /2 width=1/ /3 width=1/ -] +#G #L #V #T #H [ elim H -H /3 width=1/ ] +elim (cix_inv_ib2 … H) -H /2 width=1/ /3 width=1/ qed-. -lemma cix_inv_appl: ∀h,g,L,V,T. ⦃G, L⦄ ⊢ 𝐈[h, g]⦃ⓐV.T⦄ → +lemma cix_inv_appl: ∀h,g,G,L,V,T. ⦃G, L⦄ ⊢ 𝐈[h, g]⦃ⓐV.T⦄ → ∧∧ ⦃G, L⦄ ⊢ 𝐈[h, g]⦃V⦄ & ⦃G, L⦄ ⊢ 𝐈[h, g]⦃T⦄ & 𝐒⦃T⦄. -#h #g #L #V #T #HVT @and3_intro /3 width=1/ +#h #g #G #L #V #T #HVT @and3_intro /3 width=1/ generalize in match HVT; -HVT elim T -T // * // #a * #U #T #_ #_ #H elim H -H /2 width=1/ qed-. -lemma cix_inv_flat: ∀h,g,I,L,V,T. ⦃G, L⦄ ⊢ 𝐈[h, g]⦃ⓕ{I}V.T⦄ → +lemma cix_inv_flat: ∀h,g,I,G,L,V,T. ⦃G, L⦄ ⊢ 𝐈[h, g]⦃ⓕ{I}V.T⦄ → ∧∧ ⦃G, L⦄ ⊢ 𝐈[h, g]⦃V⦄ & ⦃G, L⦄ ⊢ 𝐈[h, g]⦃T⦄ & 𝐒⦃T⦄ & I = Appl. -#h #g * #L #V #T #H +#h #g * #G #L #V #T #H [ elim (cix_inv_appl … H) -H /2 width=1/ | elim (cix_inv_ri2 … H) -H /2 width=1/ ] @@ -63,31 +63,31 @@ qed-. (* Basic forward lemmas *****************************************************) -lemma cix_inv_cir: ∀h,g,L,T. ⦃G, L⦄ ⊢ 𝐈[h, g]⦃T⦄ → ⦃G, L⦄ ⊢ 𝐈⦃T⦄. +lemma cix_inv_cir: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ 𝐈[h, g]⦃T⦄ → ⦃G, L⦄ ⊢ 𝐈⦃T⦄. /3 width=1/ qed-. (* Basic properties *********************************************************) -lemma cix_sort: ∀h,g,L,k. deg h g k 0 → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃⋆k⦄. -#h #g #L #k #Hk #H elim (crx_inv_sort … H) -L #l #Hkl +lemma cix_sort: ∀h,g,G,L,k. deg h g k 0 → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃⋆k⦄. +#h #g #G #L #k #Hk #H elim (crx_inv_sort … H) -L #l #Hkl lapply (deg_mono … Hk Hkl) -h -k