X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground%2Frelocation%2Frtmap_coafter.ma;h=aede6b3a31512f86a2e2afbf0adec99cd4e0bc1d;hp=e74866dac69eb1b77631eb88233dab2c52cdb9c2;hb=8fdf1af656038d0245eba64ff2531bbe94ce0e9e;hpb=77c9255de3c5f7780aeacd745703a1cc76328a68 diff --git a/matita/matita/contribs/lambdadelta/ground/relocation/rtmap_coafter.ma b/matita/matita/contribs/lambdadelta/ground/relocation/rtmap_coafter.ma index e74866dac..aede6b3a3 100644 --- a/matita/matita/contribs/lambdadelta/ground/relocation/rtmap_coafter.ma +++ b/matita/matita/contribs/lambdadelta/ground/relocation/rtmap_coafter.ma @@ -14,6 +14,7 @@ include "ground/notation/relations/rcoafter_3.ma". include "ground/relocation/rtmap_sor.ma". +include "ground/relocation/rtmap_nat.ma". include "ground/relocation/rtmap_after.ma". (* RELOCATION MAP ***********************************************************) @@ -238,21 +239,21 @@ lemma coafter_mono_eq: ∀f1,f2,f. f1 ~⊚ f2 ≘ f → ∀g1,g2,g. g1 ~⊚ g2 (* Forward lemmas with pushs ************************************************) -lemma coafter_fwd_pushs: ∀j,i,g2,f1,g. g2 ~⊚ ⫯*[i]f1 ≘ g → @❪i, g2❫ ≘ j → - ∃f. ⫯*[j] f = g. -#j elim j -j -[ #i #g2 #f1 #g #Hg #H - elim (at_inv_xxp … H) -H [|*: // ] #f2 #H1 #H2 destruct - /2 width=2 by ex_intro/ -| #j #IH * [| #i ] #g2 #f1 #g #Hg #H - [ elim (at_inv_pxn … H) -H [|*: // ] #f2 #Hij #H destruct +lemma coafter_fwd_pushs: ∀k,l,g2,f1,g. g2 ~⊚ ⫯*[l]f1 ≘ g → @↑❪l, g2❫ ≘ k → + ∃∃f. ⫱*[k]g2 ~⊚ f1 ≘ f & ⫯*[k] f = g. +#k @(nat_ind_succ … k) -k +[ #l #g2 #f1 #g #Hg #H + elim (rm_nat_inv_xxp … H) -H [|*: // ] #f2 #H1 #H2 destruct + /2 width=3 by ex2_intro/ +| #k #IH * [| #l ] #g2 #f1 #g #Hg #H + [ elim (rm_nat_inv_pxn … H) -H [|*: // ] #f2 #Hlk #H destruct elim (coafter_inv_nxx … Hg) -Hg [|*: // ] #f #Hf #H destruct - elim (IH … Hf Hij) -f1 -f2 -IH /2 width=2 by ex_intro/ - | elim (at_inv_nxn … H) -H [1,4: * |*: // ] #f2 #Hij #H destruct + elim (IH … Hf Hlk) -IH -Hf -Hlk /2 width=3 by ex2_intro/ + | elim (rm_nat_inv_nxn … H) -H [1,4: * |*: // ] #f2 #Hlk #H destruct [ elim (coafter_inv_ppx … Hg) -Hg [|*: // ] #f #Hf #H destruct - elim (IH … Hf Hij) -f1 -f2 -i /2 width=2 by ex_intro/ + elim (IH … Hf Hlk) -IH -Hf -Hlk /2 width=3 by ex2_intro/ | elim (coafter_inv_nxx … Hg) -Hg [|*: // ] #f #Hf #H destruct - elim (IH … Hf Hij) -f1 -f2 -i /2 width=2 by ex_intro/ + elim (IH … Hf Hlk) -IH -Hf -Hlk /2 width=3 by ex2_intro/ ] ] ] @@ -278,14 +279,14 @@ qed-. (* Properties with iterated tail ********************************************) -lemma coafter_tls: ∀j,i,f1,f2,f. @❪i, f1❫ ≘ j → - f1 ~⊚ f2 ≘ f → ⫱*[j]f1 ~⊚ ⫱*[i]f2 ≘ ⫱*[j]f. -#j elim j -j [ #i | #j #IH * [| #i ] ] #f1 #f2 #f #Hf1 #Hf -[ elim (at_inv_xxp … Hf1) -Hf1 [ |*: // ] #g1 #Hg1 #H1 destruct // -| elim (at_inv_pxn … Hf1) -Hf1 [ |*: // ] #g1 #Hg1 #H1 +lemma coafter_tls: ∀l2,l1,f1,f2,f. @↑❪l1, f1❫ ≘ l2 → + f1 ~⊚ f2 ≘ f → ⫱*[l2]f1 ~⊚ ⫱*[l1]f2 ≘ ⫱*[l2]f. +#l2 @(nat_ind_succ … l2) -l2 [ #l1 | #l2 #IH * [| #l1 ] ] #f1 #f2 #f #Hf1 #Hf +[ elim (rm_nat_inv_xxp … Hf1) -Hf1 [ |*: // ] #g1 #Hg1 #H1 destruct // +| elim (rm_nat_inv_pxn … Hf1) -Hf1 [ |*: // ] #g1 #Hg1 #H1 elim (coafter_inv_nxx … Hf … H1) -Hf #g #Hg #H0 destruct lapply (IH … Hg1 Hg) -IH -Hg1 -Hg // -| elim (at_inv_nxn … Hf1) -Hf1 [1,4: * |*: // ] #g1 #Hg1 #H1 +| elim (rm_nat_inv_xxn … Hf1) -Hf1 [1,3: * |*: // ] #g1 [ #k1 ] #Hg1 [ #H ] #H1 [ elim (coafter_inv_pxx … Hf … H1) -Hf * #g2 #g #Hg #H2 #H0 destruct lapply (IH … Hg1 Hg) -IH -Hg1 -Hg #H // | elim (coafter_inv_nxx … Hf … H1) -Hf #g #Hg #H0 destruct @@ -294,24 +295,25 @@ lemma coafter_tls: ∀j,i,f1,f2,f. @❪i, f1❫ ≘ j → ] qed. -lemma coafter_tls_O: ∀n,f1,f2,f. @❪0, f1❫ ≘ n → - f1 ~⊚ f2 ≘ f → ⫱*[n]f1 ~⊚ f2 ≘ ⫱*[n]f. +lemma coafter_tls_O: ∀k,f1,f2,f. @↑❪𝟎, f1❫ ≘ k → + f1 ~⊚ f2 ≘ f → ⫱*[k]f1 ~⊚ f2 ≘ ⫱*[k]f. /2 width=1 by coafter_tls/ qed. +(* Note: this does not require ↑ first and second j *) lemma coafter_tls_succ: ∀g2,g1,g. g2 ~⊚ g1 ≘ g → - ∀n. @❪0, g2❫ ≘ n → ⫱*[↑n]g2 ~⊚ ⫱g1 ≘ ⫱*[↑n]g. -#g2 #g1 #g #Hg #n #Hg2 + ∀j. @❪𝟏, g2❫ ≘ j → ⫱*[j]g2 ~⊚ ⫱g1 ≘ ⫱*[j]g. +#g2 #g1 #g #Hg #j #Hg2 +lapply (rm_nat_pred_bi … Hg2) -Hg2 #Hg2 lapply (coafter_tls … Hg2 … Hg) -Hg #Hg lapply (at_pxx_tls … Hg2) -Hg2 #H elim (at_inv_pxp … H) -H [ |*: // ] #f2 #H2 -elim (coafter_inv_pxx … Hg … H2) -Hg * #f1 #f #Hf #H1 #H0 destruct -(npsucc_pred j)