From: Ferruccio Guidi Date: Sat, 10 May 2014 18:35:27 +0000 (+0000) Subject: advances on cofrees X-Git-Tag: make_still_working~925 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=a0d25627e80a3a2fe68da954b68f6d541c6dbc34;p=helm.git advances on cofrees --- diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/llpx_sn_alt2.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/llpx_sn_alt2.etc index bdbc32e88..888f83ef4 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/llpx_sn_alt2.etc +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/llpx_sn_alt2.etc @@ -15,6 +15,104 @@ include "basic_2/substitution/cofrees_lift.ma". include "basic_2/substitution/llpx_sn_alt1.ma". +lemma le_plus_xSy_O_false: ∀x,y. x + S y ≤ 0 → ⊥. +#x #y #H lapply (le_n_O_to_eq … H) -H minus_plus_plus_l + #Hd #He lapply (le_plus_to_le_r … Hd) -Hd + #Hd >IHL12 // -L2 >plus_minus /2 width=3 by transitive_le/ +] +qed-. + +lemma ldrop_fwd_length_le_ge: ∀L1,L2,d,e,s. ⇩[s, d, e] L1 ≡ L2 → d ≤ |L1| → |L1| - d ≤ e → |L2| = d. +#L1 #L2 #d #e #s #H elim H -L1 -L2 -d -e normalize +[ /2 width=1 by le_n_O_to_eq/ +| #I #L #V #_ minus_plus_plus_l + #L #HL1 #HL2 elim (lt_or_ge (|L1|) (e2-e1)) #H0 + [ elim (ldrop_inv_O1_gt … HL1 H0) -HL1 #H1 #H2 destruct + elim (ldrop_inv_atom1 … HL2) -HL2 #H #_ destruct + @(ex2_intro … (⋆)) [ @ldrop_O1_ge normalize // ] + @ldrop_atom #H destruct + | elim (ldrop_O1_pair … HL1 H0 I V) -HL1 -H0 /3 width=5 by ldrop_drop, ex2_intro/ + ] + ] +| #I #L1 #L2 #V1 #V2 #d #e2 #_ #HV21 #IHL12 #e1 #He12 elim (IHL12 … He12) -IHL12 + #L #HL1 #HL2 elim (lift_split … HV21 d e1) -HV21 /3 width=5 by ldrop_skip, ex2_intro/ +] +qed-. + (* LAZY SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS ****) (* alternative definition of llpx_sn (not recursive) *) @@ -34,6 +132,153 @@ lemma cpy_inv_nlift2_be: ∀G,L,U1,U2,d. ⦃G, L⦄ ⊢ U1 ▶[d, ∞] U2 → elim (cpy_inv_lift1_be … HU12 … HLK … HTU1) /2 width=2 by/ qed-. +lemma cpy_inv_nlift2_ge: ∀G,L,U1,U2,d,e. ⦃G, L⦄ ⊢ U1 ▶[d, e] U2 → + ∀i. d ≤ yinj i → (* yinj i + yinj 1 ≤ d + e → *) + (∀T2. ⇧[i, 1] T2 ≡ U2 → ⊥) → + ∃∃j. d ≤ yinj j & j ≤ i & (∀T1. ⇧[j, 1] T1 ≡ U1 → ⊥). +#G #L #U1 #U2 #d #e #H elim H -G -L -U1 -U2 -d -e +[ #I #G #L #d #e #i #Hdi (* #Hide *) #H @(ex3_intro … i) /2 width=2 by/ +| #I #G #L #K #V #W #j #d #e #Hdj #Hjde #HLK #HVW #i #Hdi (* #Hide *) #HnW + elim (le_or_ge i j) #Hij [2: @(ex3_intro … j) /2 width=7 by lift_inv_lref2_be/ ] + elim (lift_split … HVW i j) -HVW // + #X #_ #H elim HnW -HnW // +| #a #I #G #L #W1 #W2 #U1 #U2 #d #e #_ #_ #IHW12 #IHU12 #i #Hdi #H elim (nlift_inv_bind … H) -H + [ #HnW2 elim (IHW12 … HnW2) -IHW12 -HnW2 -IHU12 // + #j #Hdj #Hji #HnW1 @(ex3_intro … j) /3 width=9 by nlift_bind_sn/ + | #HnU2 elim (IHU12 … HnU2) -IHU12 -HnU2 -IHW12 /2 width=1 by yle_succ/ + #j #Hdj #Hji + >(plus_minus_m_m j 1) in ⊢ (%→?); [2: /3 width=3 by yle_trans, yle_inv_inj/ ] + #HnW1 @(ex3_intro … (j-1)) /3 width=9 by nlift_bind_dx, yle_pred, monotonic_pred/ + ] +| #I #G #L #W1 #W2 #U1 #U2 #d #e #_ #_ #IHW12 #IHU12 #i #Hdi #H elim (nlift_inv_flat … H) -H + [ #HnW2 elim (IHW12 … HnW2) -IHW12 -HnW2 -IHU12 // + #j #Hdj #Hji #HnW1 @(ex3_intro … j) /3 width=8 by nlift_flat_sn/ + | #HnU2 elim (IHU12 … HnU2) -IHU12 -HnU2 -IHW12 // + #j #Hdj #Hji #HnW1 @(ex3_intro … j) /3 width=8 by nlift_flat_dx/ + ] +] +qed-. + +axiom frees_fwd_nlift_ge: ∀L,U,d,i. (L ⊢ i ~ϵ 𝐅*[d]⦃U⦄ → ⊥) → d ≤ yinj i → + ∃∃j. d ≤ yinj j & j ≤ i & (∀T. ⇧[j, 1] T ≡ U → ⊥). + +(* +lemma frees_ind_nlift: ∀L,d. ∀R:relation2 term nat. + (∀U1,i. d ≤ yinj i → (∀T1. ⇧[i, 1] T1 ≡ U1 → ⊥) → R U1 i) → + (∀U1,U2,i,j. d ≤ yinj j → j ≤ i → ⦃⋆, L⦄ ⊢ U1 ▶[d, ∞] U2 → (∀T2. ⇧[i, 1] T2 ≡ U2 → ⊥) → R U2 i → R U1 j) → + ∀U,i. d ≤ yinj i → (L ⊢ i ~ϵ 𝐅*[d]⦃U⦄ → ⊥) → R U i. +#L #d #R #IH1 #IH2 #U1 #i #Hdi #H @(frees_ind … H) -U1 /3 width=4 by/ +#U1 #U2 #HU12 #HnU2 #HU2 @(IH2 … HU12 … HU2) + +qed-. +*)(* +lemma frees_fwd_nlift: ∀L,d. ∀R:relation2 term nat. ( + ∀U1,j. (∀T1. ⇧[j, 1] T1 ≡ U1 → ⊥) ∨ + (∃∃U2,i. d ≤ yinj j → j < i & (L ⊢ j ~ϵ 𝐅*[d]⦃U1⦄ → ⊥) & ⦃⋆, L⦄ ⊢ U1 ▶[d, ∞] U2 & R U1 i + ) → + d ≤ yinj j → R U1 j + ) → + ∀U,i. d ≤ yinj i → (L ⊢ i ~ϵ 𝐅*[d]⦃U⦄ → ⊥) → R U i. +#L #d #R #IHR #U1 #j #Hdj #H elim (frees_inv_gen … H) -H +#U2 #H generalize in match Hdj; -Hdj generalize in match j; -j @(cpys_ind … H) -U2 +[ #j #Hdj #HnU1 @IHR -IHR /3 width=2 by or_introl/ +| #U0 #U2 #HU10 #HU02 #IHU10 #j #Hdj #HnU2 elim (cpy_inv_nlift2_ge … HU02 … Hdj HnU2) -HU02 -HnU2 + #i #Hdi #Hij #HnU0 lapply (IHU10 … HnU0) // -IHU10 + #Hi @IHR -IHR // -Hdj @or_intror + +lemma frees_fwd_nlift: ∀L,U,d,i. d ≤ yinj i → (L ⊢ i ~ϵ 𝐅*[d]⦃U⦄ → ⊥) → + ∃∃j. d ≤ yinj j & j ≤ i & (∀T. ⇧[j, 1] T ≡ U → ⊥). +#L #U1 #d #i #Hdi #H +#U2 #H #HnU2 @(cpys_ind_dx … H) -U1 [ @(ex3_intro … i) /2 width=2 by/ ] -Hdi -HnU2 +#U1 #U0 #HU10 #_ * #j #Hdj #Hji #HnU0 elim (cpy_inv_nlift2_ge … HU10 … Hdj HnU0) -U0 -Hdj +/3 width=5 by transitive_le, ex3_intro/ +qed-. +*) + +theorem llpx_sn_llpx_sn_alt2: ∀R,L1,L2,T,d. llpx_sn R d T L1 L2 → llpx_sn_alt2 R d T L1 L2. +#R #L1 #L2 #U1 #d #H elim (llpx_sn_inv_alt1 … H) -H +#HL12 #IH @conj // -HL12 +#I1 #I2 #K1 #K2 #V1 #V2 #i #Hdi #H #HLK1 #HLK2 +elim (frees_fwd_nlift_ge … H Hdi) -H -Hdi #j #Hdj #Hji #HnU1 +lapply (ldrop_fwd_length_lt2 … HLK1) #HL1 +lapply (ldrop_fwd_length_lt2 … HLK2) #HL2 +lapply (le_to_lt_to_lt … Hji HL1) -HL1 #HL1 +lapply (le_to_lt_to_lt … Hji HL2) -HL2 #HL2 +elim (ldrop_O1_lt L1 j) // #Z1 #Y1 #X1 #HLY1 +elim (ldrop_O1_lt L2 j) // #Z2 #Y2 #X2 #HLY2 + + + + +generalize in match V2; -V2 generalize in match V1; -V1 +generalize in match K2; -K2 generalize in match K1; -K1 +generalize in match I2; -I2 generalize in match I1; -I1 +generalize in match IH; -IH +@(frees_ind_nlift … Hdi H) -U1 -i +[ #U1 #i #Hdi #HnU1 #IH #I1 #I2 #K1 #K2 #V1 #V2 #HLK1 #HLK2 elim (IH … HnU1 HLK1 HLK2) -IH -HnU1 -HLK1 -HLK2 /2 width=1 by conj/ +| #U1 #U2 #i #j #Hdj #Hji #HU12 #HnU2 #IHU12 #IH #I1 #I2 #K1 #K2 #V1 #V2 #HLK1 #HLK2 +(* +*) + @(IHU12) … HLK1 HLK2) + + @(IHU12 … HLK1 HLK2) -IHU02 -I1 -I2 -K1 -K2 -V1 -V2 + #I1 #I2 #K1 #K2 #V1 #V2 #j #Hdj #HnU0 #HLK1 #HLK2 @(IH … HLK1 HLK2) -IH -HLK1 -HLK2 // + + + + elim (frees_fwd_nlift … HnU1) // -HnU1 -Hdi +#j #Hdj #Hji #HnU1 +lapply (ldrop_fwd_length_lt2 … HLK1) #HL1 +lapply (ldrop_fwd_length_lt2 … HLK2) #HL2 +lapply (le_to_lt_to_lt … Hji HL1) -HL1 #HL1 +lapply (le_to_lt_to_lt … Hji HL2) -HL2 #HL2 +elim (ldrop_O1_lt L1 j) // #Z1 #Y1 #X1 #HLY1 +elim (ldrop_O1_lt L2 j) // #Z2 #Y2 #X2 #HLY2 +elim (IH … HnU1 HLY1 HLY2) // #H #HX12 #HY12 destruct + + + + + + + + + + + + + + + + + + + + + + +theorem llpx_sn_llpx_sn_alt2: ∀R,L1,L2,T,d. llpx_sn R d T L1 L2 → llpx_sn_alt2 R d T L1 L2. +#R #L1 #L2 #U1 #d #H @(llpx_sn_ind_alt1 … H) -L1 -L2 -U1 -d +#L1 #L2 #U1 #d #HL12 #IH @conj // -HL12 +#I1 #I2 #K1 #K2 #V1 #V2 #i #Hdi #HnU1 #HLK1 #HLK2 elim (frees_inv_gen … HnU1) -HnU1 +#U2 #H generalize in match IH; -IH @(cpys_ind_dx … H) -U1 +[ #IH #HnU2 elim (IH … HnU2 … HLK1 HLK2) -L1 -L2 -U2 /2 width=1 by conj/ +| #U1 #U0 #HU10 #_ #IHU02 #IH #HnU2 @IHU02 /2 width=2 by/ -I1 -I2 -K1 -K2 -V1 -V2 -U2 -i + #I1 #I2 #K1 #K2 #V1 #V2 #i #Hdi #HnU0 #HLK1 #HLK2 @(IH … HLK1 HLK2) -IH // -R -I2 -L2 -K2 -V2 + @(cpy_inv_nlift2_be … HU10) /2 width=3 by/ + +theorem llpx_sn_llpx_sn_alt2: ∀R,L1,L2,T2,d. llpx_sn R d T2 L1 L2 → + ∀T1. ⦃⋆, L1⦄ ⊢ T1 ▶*[d, ∞] T2 → llpx_sn_alt2 R d T1 L1 L2. +#R #L1 #L2 #U2 #d #H elim (llpx_sn_inv_alt1 … H) -H +#HL12 #IH #U1 #HU12 @conj // -HL12 +#I1 #I2 #K1 #K2 #V1 #V2 #i #Hdi #HnU1 #HLK1 #HLK2 elim (frees_inv_gen … HnU1) -HnU1 +#U2 #H generalize in match IH; -IH @(cpys_ind_dx … H) -U1 +[ #IH #HnU2 elim (IH … HnU2 … HLK1 HLK2) -L1 -L2 -U2 /2 width=1 by conj/ +| #U1 #U0 #HU10 #_ #IHU02 #IH #HnU2 @IHU02 /2 width=2 by/ -I1 -I2 -K1 -K2 -V1 -V2 -U2 -i + #I1 #I2 #K1 #K2 #V1 #V2 #i #Hdi #HnU0 #HLK1 #HLK2 @(IH … HLK1 HLK2) -IH // -R -I2 -L2 -K2 -V2 + @(cpy_inv_nlift2_be … HU10) /2 width=3 by/ + + theorem llpx_sn_llpx_sn_alt2: ∀R,L1,L2,T,d. llpx_sn R d T L1 L2 → llpx_sn_alt2 R d T L1 L2. #R #L1 #L2 #U1 #d #H elim (llpx_sn_inv_alt1 … H) -H #HL12 #IH @conj // -HL12 diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cofrees.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cofrees.ma index 0b311d0d2..30ecd819a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cofrees.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/cofrees.ma @@ -71,9 +71,19 @@ lemma cofrees_flat: ∀L,V,d,i. L ⊢ i ~ϵ 𝐅*[d]⦃V⦄ → ∀T. L ⊢ i ~ elim (HW1 … HW12) elim (HU1 … HU12) -W1 -U1 /3 width=2 by lift_flat, ex_intro/ qed. +lemma cofrees_cpy_trans: ∀L,U1,U2,d. ⦃⋆, L⦄ ⊢ U1 ▶[d, ∞] U2 → + ∀i. L ⊢ i ~ϵ 𝐅*[d]⦃U1⦄ → L ⊢ i ~ϵ 𝐅*[d]⦃U2⦄. +/3 width=3 by cpys_strap2/ qed-. + axiom cofrees_dec: ∀L,T,d,i. Decidable (L ⊢ i ~ϵ 𝐅*[d]⦃T⦄). -(* Negated inversion lemmas *************************************************) +(* Basic negated properties *************************************************) + +lemma frees_cpy_div: ∀L,U1,U2,d. ⦃⋆, L⦄ ⊢ U1 ▶[d, ∞] U2 → + ∀i. (L ⊢ i ~ϵ 𝐅*[d]⦃U2⦄ → ⊥) → (L ⊢ i ~ϵ 𝐅*[d]⦃U1⦄ → ⊥). +/3 width=7 by cofrees_cpy_trans/ qed-. + +(* Basic negated inversion lemmas *******************************************) lemma frees_inv_bind: ∀a,I,L,V,T,d,i. (L ⊢ i ~ϵ 𝐅*[d]⦃ⓑ{a,I}V.T⦄ → ⊥) → (L ⊢ i ~ϵ 𝐅*[d]⦃V⦄ → ⊥) ∨ (L.ⓑ{I}V ⊢ i+1 ~ϵ 𝐅*[⫯d]⦃T⦄ → ⊥). diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cofrees_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cofrees_lift.ma index 30bcd9ce9..b8e4542a7 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cofrees_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/cofrees_lift.ma @@ -95,3 +95,11 @@ lemma frees_inv_gen: ∀L,U,d,i. (L ⊢ i ~ϵ 𝐅*[d]⦃U⦄ → ⊥) → /4 width=9 by cpys_flat, nlift_flat_dx, nlift_flat_sn, ex2_intro/ ] qed-. + +lemma frees_ind: ∀L,d,i. ∀R:predicate term. + (∀U1. (∀T1. ⇧[i, 1] T1 ≡ U1 → ⊥) → R U1) → + (∀U1,U2. ⦃⋆, L⦄ ⊢ U1 ▶[d, ∞] U2 → (L ⊢ i ~ϵ 𝐅*[d]⦃U2⦄ → ⊥) → R U2 → R U1) → + ∀U. (L ⊢ i ~ϵ 𝐅*[d]⦃U⦄ → ⊥) → R U. +#L #d #i #R #IH1 #IH2 #U1 #H elim (frees_inv_gen … H) -H +#U2 #H #HnU2 @(cpys_ind_dx … H) -U1 /4 width=8 by cofrees_inv_gen/ +qed-.