From 926796df5884453d8f0cf9f294d7776d469ef45b Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi <ferruccio.guidi@unibo.it> Date: Mon, 6 Jun 2016 20:20:44 +0000 Subject: [PATCH] work in progress on frees_drops --- .../lambdadelta/basic_2/relocation/drops.ma | 12 ++ .../lambdadelta/basic_2/relocation/lifts.ma | 30 +++- .../lambdadelta/basic_2/static/frees_drops.ma | 163 +++++++++++++++++- .../ground_2/relocation/rtmap_at.ma | 5 + .../ground_2/relocation/rtmap_coafter.ma | 77 ++++++++- .../ground_2/relocation/rtmap_isfin.ma | 8 +- 6 files changed, 276 insertions(+), 19 deletions(-) diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma index 8d22bcf00..9a4cfca83 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma @@ -282,6 +282,18 @@ lemma drops_inv_TF: âf,I,L,K,V. â¬*[â», f] L â¡ K.â{I}V â ðâ¦f⦠(* Advanced inversion lemmas ************************************************) +lemma drops_inv_atom2: âb,L,f. â¬*[b,f] L â¡ â â + âân,f1. â¬*[b,ðâ´nâµ] L â¡ â & ðâ´nâµ â f1 â¡ f. +#b #L elim L -L +[ /3 width=4 by drops_atom, after_isid_sn, ex2_2_intro/ +| #L #I #V #IH #f #H elim (pn_split f) * #g #H0 destruct + [ elim (drops_inv_skip1 ⦠H) -H #K #W #_ #_ #H destruct + | lapply (drops_inv_drop1 ⦠H) -H #HL + elim (IH ⦠HL) -IH -HL /3 width=8 by drops_drop, after_next, ex2_2_intro/ + ] +] +qed-. + (* Basic_2A1: includes: drop_inv_gen *) lemma drops_inv_gen: âb,f,I,L,K,V. â¬*[b, f] L â¡ K.â{I}V â ðâ¦f⦠â â¬*[â, f] L â¡ K.â{I}V. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts.ma index a621b819f..be42c0ebf 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts.ma @@ -210,6 +210,28 @@ lemma lifts_inv_flat2: âf:rtmap. âI,V2,T2,X. â¬*[f] X â¡ â{I}V2.T2 â (* Advanced inversion lemmas ************************************************) +lemma lifts_inv_atom1: âf,I,Y. â¬*[f] âª{I} â¡ Y â + â¨â¨ ââs. I = Sort s & Y = âs + | ââi,j. @â¦i, f⦠⡠j & I = LRef i & Y = #j + | ââl. I = GRef l & Y = §l. +#f * #n #Y #H +[ lapply (lifts_inv_sort1 ⦠H) +| elim (lifts_inv_lref1 ⦠H) +| lapply (lifts_inv_gref1 ⦠H) +] -H /3 width=5 by or3_intro0, or3_intro1, or3_intro2, ex3_2_intro, ex2_intro/ +qed-. + +lemma lifts_inv_atom2: âf,I,X. â¬*[f] X â¡ âª{I} â + â¨â¨ ââs. X = âs & I = Sort s + | ââi,j. @â¦i, f⦠⡠j & X = #i & I = LRef j + | ââl. X = §l & I = GRef l. +#f * #n #X #H +[ lapply (lifts_inv_sort2 ⦠H) +| elim (lifts_inv_lref2 ⦠H) +| lapply (lifts_inv_gref2 ⦠H) +] -H /3 width=5 by or3_intro0, or3_intro1, or3_intro2, ex3_2_intro, ex2_intro/ +qed-. + (* Basic_2A1: includes: lift_inv_pair_xy_x *) lemma lifts_inv_pair_xy_x: âf,I,V,T. â¬*[f] â¡{I}V.T â¡ V â â¥. #f #J #V elim V -V @@ -249,14 +271,6 @@ qed-. (* Basic forward lemmas *****************************************************) -lemma lifts_fwd_atom2: âf,I,X. â¬*[f] X â¡ âª{I} â âJ. X = âª{J}. -#f * #n #X #H -[ lapply (lifts_inv_sort2 ⦠H) -| elim (lifts_inv_lref2 ⦠H) -| lapply (lifts_inv_gref2 ⦠H) -] -H /2 width=2 by ex_intro/ -qed-. - (* Basic_2A1: includes: lift_inv_O2 *) lemma lifts_fwd_isid: âf,T1,T2. â¬*[f] T1 â¡ T2 â ðâ¦f⦠â T1 = T2. #f #T1 #T2 #H elim H -f -T1 -T2 diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma index fd4ec0ce1..aa4f34998 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma @@ -14,19 +14,19 @@ include "ground_2/relocation/rtmap_pushs.ma". include "ground_2/relocation/rtmap_coafter.ma". -include "basic_2/relocation/lifts_lifts.ma". -include "basic_2/relocation/drops.ma". +include "basic_2/relocation/drops_drops.ma". include "basic_2/static/frees.ma". (* CONTEXT-SENSITIVE FREE VARIABLES *****************************************) (* Advanced properties ******************************************************) -lemma frees_lref_atom: âL,i. â¬*[â», ðâ´iâµ] L â¡ â â L ⢠ð *â¦#i⦠⡠ðð. -#L elim L -L /2 width=1 by frees_atom/ +lemma frees_lref_atom: âb,L,i. â¬*[b, ðâ´iâµ] L â¡ â â + âf. ðâ¦f⦠â L ⢠ð *â¦#i⦠⡠f. +#b #L elim L -L /2 width=1 by frees_atom/ #L #I #V #IH * [ #H lapply (drops_fwd_isid ⦠H ?) -H // #H destruct -| /4 width=3 by frees_lref, drops_inv_drop1/ +| /5 width=3 by frees_eq_repl_back, frees_lref, drops_inv_drop1, eq_push_inv_isid/ ] qed. @@ -58,6 +58,91 @@ qed-. (* Properties with generic slicing for local environments *******************) +axiom coafter_inv_xpx: âg2,f1,g. g2 ~â âf1 â¡ g â ân. @â¦0, g2⦠⡠n â + ââf2,f. f2 ~â f1 â¡ f & ⫱*[n]g2 = âf2 & ⫱*[n]g = âf. +(* +#g2 #g1 #g #Hg #n #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 +<tls_rew_S <tls_rew_S <H2 <H0 -g2 -g -n // +qed. +*) + +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 +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 +<tls_rew_S <tls_rew_S <H2 <H0 -g2 -g -n // +qed. + +lemma frees_lifts: âb,f1,K,T. K ⢠ð *â¦T⦠⡠f1 â + âf,L. â¬*[b, f] L â¡ K â âU. â¬*[f] T â¡ U â + âf2. f ~â f1 â¡ f2 â L ⢠ð *â¦U⦠⡠f2. +#b #f1 #K #T #H lapply (frees_fwd_isfin ⦠H) elim H -f1 -K -T +[ #f1 #I #Hf1 #_ #f #L #H1 #U #H2 #f2 #H3 + lapply (coafter_isid_inv_dx ⦠H3 ⦠Hf1) -f1 #Hf2 + elim (lifts_inv_atom1 ⦠H2) -H2 * + /2 width=1 by frees_sort_gen, frees_gref_gen/ + #i #j #Hij #H #H0 destruct + elim (drops_inv_atom2 ⦠H1) -H1 #n #g #H1 #Hf + elim (after_at_fwd ⦠Hij ⦠Hf) -f #x #_ #Hj -g -i + lapply (at_inv_uni ⦠Hj) -Hj #H destruct + /3 width=8 by frees_lref_atom, drops_trans/ +| #f1 #I #K #V #s #_ #IH #Hf1 #f #L #H1 #U #H2 #f2 #H3 + lapply (isfin_fwd_push ⦠Hf1 ??) -Hf1 [3: |*: // ] #Hf1 + lapply (lifts_inv_sort1 ⦠H2) -H2 #H destruct + lapply (at_total 0 f) #H + elim (drops_split_trans ⦠H1) -H1 + [5: @(after_uni_dx ⦠H) /2 width=1 by after_isid_dx/ |2,3: skip + |4: // ] #X #HLX #HXK + lapply (drops_inv_tls_at ⦠H ⦠HXK) -HXK #HXK + elim (drops_inv_skip2 ⦠HXK) -HXK + #Y #W #HYK #HVW #H0 destruct +(* + + elim (coafter_inv_xpx ⦠H3 ??) -H3 [ |*: // ] #g2 #g #Hg #H2 #H0 + lapply (IH ⦠Hg) -IH -Hg + [1,5: // | skip + | + |6: #H +*) + + lapply (coafter_tls_succ ⦠H3 ??) -H3 [3: |*: // ] #H3 + lapply (IH ⦠HYK ⦠H3) -IH -H3 -HYK + [1,3: // | skip ] + #H lapply (frees_sort ⦠H) + + ] + + + elim (coafter_inv_xxp ⦠H3) -H3 [1,3: * |*: // ] + [ #g #g1 #Hf2 #H #H0 destruct + elim (drops_inv_skip1 ⦠H1) -H1 #K #V #HLK #_ #H destruct + | #g #Hf2 #H destruct + lapply (drops_inv_drop1 ⦠H1) -H1 + ] /3 width=4 by frees_sort/ + +| +| +| +| #f1V #f1T #f1 #p #I #K #V #T #_ #_ #H1f1 #IHV #IHT #H2f1 #f #L #H1 #Y #H2 #f2 #H3 + elim (sor_inv_isfin3 ⦠H1f1) // #Hf1V #H + lapply (isfin_inv_tl ⦠H) -H + elim (lifts_inv_bind1 ⦠H2) -H2 #W #U #HVW #HTU #H destruct + elim (coafter_sor ⦠H3 ⦠H1f1) /2 width=5 by coafter_isfin2_fwd/ -H3 -H1f1 #f2V #f2T #Hf2V #H + elim (coafter_inv_tl1 ⦠H) -H /4 width=5 by frees_bind, drops_skip/ +| #f1V #f1T #f1 #I #K #V #T #_ #_ #H1f1 #IHV #IHT #H2f1 #f #L #H1 #Y #H2 #f2 #H3 + elim (sor_inv_isfin3 ⦠H1f1) // + elim (lifts_inv_flat1 ⦠H2) -H2 #W #U #HVW #HTU #H destruct + elim (coafter_sor ⦠H3 ⦠H1f1) + /3 width=5 by coafter_isfin2_fwd, frees_flat/ +] + (* Inversion lemmas with generic slicing for local environments *************) lemma frees_inv_lifts: âb,f2,L,U. L ⢠ð *â¦U⦠⡠f2 â @@ -67,8 +152,7 @@ lemma frees_inv_lifts: âb,f2,L,U. L ⢠ð *â¦U⦠⡠f2 â [ #f2 #I #Hf2 #_ #f #K #H1 #T #H2 #f1 #H3 lapply (coafter_fwd_isid2 ⦠H3 ⦠Hf2) -H3 // -Hf2 #Hf1 elim (drops_inv_atom1 ⦠H1) -H1 #H #_ destruct - elim (lifts_fwd_atom2 ⦠H2) -H2 - /2 width=3 by frees_atom/ + elim (lifts_inv_atom2 ⦠H2) -H2 * /2 width=3 by frees_atom/ | #f2 #I #L #W #s #_ #IH #Hf2 #f #Y #H1 #T #H2 #f1 #H3 lapply (isfin_fwd_push ⦠Hf2 ??) -Hf2 [3: |*: // ] #Hf2 lapply (lifts_inv_sort2 ⦠H2) -H2 #H destruct @@ -117,3 +201,68 @@ lemma frees_inv_lifts: âb,f2,L,U. L ⢠ð *â¦U⦠⡠f2 â elim (coafter_inv_sor ⦠H3 ⦠H1f2) -H3 -H1f2 /3 width=5 by frees_flat/ ] qed-. + +lemma frees_inv_drops: âf2,L,U. L ⢠ð *â¦U⦠⡠f2 â + âf,K. â¬*[â, f] L â¡ K â âf1. f ~â f1 â¡ f2 â + ââT. K ⢠ð *â¦T⦠⡠f1 & â¬*[f] T â¡ U. +#f2 #L #U #H lapply (frees_fwd_isfin ⦠H) elim H -f2 -L -U +[ #f2 #I #Hf2 #_ #f #K #H1 #f1 #H2 + lapply (coafter_fwd_isid2 ⦠H2 ??) -H2 // -Hf2 #Hf1 + elim (drops_inv_atom1 ⦠H1) -H1 #H #Hf destruct + /4 width=3 by frees_atom, lifts_refl, ex2_intro/ +| #f2 #I #L #W #s #_ #IH #Hf2 #f #Y #H1 #f1 #H2 + lapply (isfin_fwd_push ⦠Hf2 ??) -Hf2 [3: |*: // ] #Hf2 + elim (coafter_inv_xxp ⦠H2) -H2 [1,3: * |*: // ] + [ #g #g1 #Hf2 #H #H0 destruct + elim (drops_inv_skip1 ⦠H1) -H1 #K #V #HLK #_ #H destruct + | #g #Hf2 #H destruct + lapply (drops_inv_drop1 ⦠H1) -H1 #HLK + ] + elim (IH ⦠HLK ⦠Hf2) -L // -f2 #X #Hg1 #HX + lapply (lifts_inv_sort2 ⦠HX) -HX #H destruct + /3 width=3 by frees_sort, lifts_sort, ex2_intro/ +| #f2 #I #L #W #_ #IH #Hf2 #f #Y #H1 #f1 #H2 + lapply (isfin_inv_next ⦠Hf2 ??) -Hf2 [3: |*: // ] #Hf2 + elim (coafter_inv_xxn ⦠H2) -H2 [ |*: // ] #g #g1 #Hf2 #H0 #H destruct + elim (drops_inv_skip1 ⦠H1) -H1 #K #V #HLK #HVW #H destruct + elim (IH ⦠HLK ⦠Hf2) -L // -f2 #X #Hg1 #HX + lapply (lifts_inj ⦠HX ⦠HVW) -W #H destruct + /3 width=3 by frees_zero, lifts_lref, ex2_intro/ +| #f2 #I #L #W #j #_ #IH #Hf2 #f #Y #H1 #f1 #H2 + lapply (isfin_fwd_push ⦠Hf2 ??) -Hf2 [3: |*: // ] #Hf2 + elim (coafter_inv_xxp ⦠H2) -H2 [1,3: * |*: // ] + [ #g #g1 #Hf2 #H #H0 destruct + elim (drops_inv_skip1 ⦠H1) -H1 #K #V #HLK #_ #H destruct + | #g #Hf2 #H destruct + lapply (drops_inv_drop1 ⦠H1) -H1 #HLK + ] + elim (IH ⦠HLK ⦠Hf2) -L // -f2 #X #Hg1 #HX + elim (lifts_inv_lref2 ⦠HX) -HX #i #Hij #H destruct + /4 width=7 by frees_lref, lifts_lref, at_S1, at_next, ex2_intro/ +| #f2 #I #L #W #l #_ #IH #Hf2 #f #Y #H1 #f1 #H2 + lapply (isfin_fwd_push ⦠Hf2 ??) -Hf2 [3: |*: // ] #Hf2 + elim (coafter_inv_xxp ⦠H2) -H2 [1,3: * |*: // ] + [ #g #g1 #Hf2 #H #H0 destruct + elim (drops_inv_skip1 ⦠H1) -H1 #K #V #HLK #_ #H destruct + | #g #Hf2 #H destruct + lapply (drops_inv_drop1 ⦠H1) -H1 #HLK + ] + elim (IH ⦠HLK ⦠Hf2) -L // -f2 #X #Hg1 #HX + lapply (lifts_inv_gref2 ⦠HX) -HX #H destruct + /3 width=3 by frees_gref, lifts_gref, ex2_intro/ +| #f2W #f2U #f2 #p #I #L #W #U #_ #_ #H1f2 #IHW #IHU #H2f2 #f #K #H1 #f1 #H2 + elim (sor_inv_isfin3 ⦠H1f2) // #H1f2W #H + lapply (isfin_inv_tl ⦠H) -H #H1f2U + elim (coafter_inv_sor ⦠H2 ⦠H1f2) -H2 -H1f2 // #f1W #f1U #H2f2W #H #Hf1 + elim (coafter_inv_tl0 ⦠H) -H #g1 #H2f2U #H destruct + elim (IHW ⦠H1 ⦠H2f2W) -IHW -H2f2W // -H1f2W #V #Hf1W #HVW + elim (IHU ⦠H2f2U) -IHU -H2f2U + /3 width=5 by frees_bind, drops_skip, lifts_bind, ex2_intro/ +| #f2W #f2U #f2 #I #L #W #U #_ #_ #H1f2 #IHW #IHU #H2f2 #f #K #H1 #f1 #H2 + elim (sor_inv_isfin3 ⦠H1f2) // #H1f2W #H1f2U + elim (coafter_inv_sor ⦠H2 ⦠H1f2) -H2 -H1f2 // #f1W #f1U #H2f2W #H2f2U #Hf1 + elim (IHW ⦠H1 ⦠H2f2W) -IHW -H2f2W // -H1f2W + elim (IHU ⦠H1 ⦠H2f2U) -L -H2f2U + /3 width=5 by frees_flat, lifts_flat, ex2_intro/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_at.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_at.ma index 90e73d347..4af66a1c2 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_at.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_at.ma @@ -385,3 +385,8 @@ theorem at_div_id_sn: âf. H_at_div ðð f f ðð. lemma at_uni: ân,i. @â¦i,ðâ´nâµâ¦ â¡ n+i. #n elim n -n /2 width=5 by at_next/ qed. + +(* Inversion lemmas with uniform relocations ********************************) + +lemma at_inv_uni: ân,i,j. @â¦i,ðâ´nâµâ¦ â¡ j â j = n+i. +/2 width=4 by at_mono/ qed-. diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_coafter.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_coafter.ma index c58db6ec0..5ac0c06f7 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_coafter.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_coafter.ma @@ -37,6 +37,9 @@ definition H_coafter_inj: predicate rtmap â definition H_coafter_fwd_isid2: predicate rtmap â λf1. âf2,f. f1 ~â f2 â¡ f â ðâ¦f1⦠â ðâ¦f⦠â ðâ¦f2â¦. +definition H_coafter_isfin2_fwd: predicate rtmap â + λf1. âf2. ð â¦f2⦠â ðâ¦f1⦠â âf. f1 ~â f2 â¡ f â ð â¦fâ¦. + (* Basic inversion lemmas ***************************************************) lemma coafter_inv_ppx: âg1,g2,g. g1 ~â g2 â¡ g â âf1,f2. âf1 = g1 â âf2 = g2 â @@ -277,11 +280,19 @@ lemma coafter_mono_eq: âf1,f2,f. f1 ~â f2 â¡ f â âg1,g2,g. g1 ~â g2 (* Inversion lemmas with tail ***********************************************) +lemma coafter_inv_tl1: âg2,g1,g. g2 ~â ⫱g1 â¡ g â + ââf. âg2 ~â g1 â¡ f & ⫱f = g. +#g2 #g1 #g elim (pn_split g1) * #f1 #H1 #H destruct +[ /3 width=7 by coafter_refl, ex2_intro/ +| @(ex2_intro ⦠(⫯g)) /2 width=7 by coafter_push/ (**) (* full auto fails *) +] +qed-. + lemma coafter_inv_tl0: âg2,g1,g. g2 ~â g1 ⡠⫱g â ââf1. âg2 ~â f1 â¡ g & ⫱f1 = g1. -#g1 #g2 #g elim (pn_split g) * #f #H0 #H destruct +#g2 #g1 #g elim (pn_split g) * #f #H0 #H destruct [ /3 width=7 by coafter_refl, ex2_intro/ -| @(ex2_intro ⦠(⫯g2)) /2 width=7 by coafter_push/ (**) (* full auto fails *) +| @(ex2_intro ⦠(⫯g1)) /2 width=7 by coafter_push/ (**) (* full auto fails *) ] qed-. @@ -586,6 +597,36 @@ lemma coafter_fwd_isid2: âf1. H_coafter_fwd_isid2 f1. /3 width=7 by coafter_fwd_isid2_aux, coafter_fwd_isid2_O_aux/ qed-. +fact coafter_isfin2_fwd_O_aux: âf1. @â¦0, f1⦠⡠0 â + H_coafter_isfin2_fwd f1. +#f1 #Hf1 #f2 #H +generalize in match Hf1; generalize in match f1; -f1 +@(isfin_ind ⦠H) -f2 +[ /3 width=4 by coafter_isid_inv_dx, isfin_isid/ ] +#f2 #_ #IH #f1 #H #Hf1 #f #Hf +elim (at_inv_pxp ⦠H) -H [ |*: // ] #g1 #H1 +lapply (istot_inv_push ⦠Hf1 ⦠H1) -Hf1 #Hg1 +elim (Hg1 0) #n #Hn +[ elim (coafter_inv_ppx ⦠Hf) | elim (coafter_inv_pnx ⦠Hf) +] -Hf [1,6: |*: // ] #g #Hg #H0 destruct +/5 width=6 by isfin_next, isfin_push, isfin_inv_tls, istot_tls, at_pxx_tls, coafter_tls/ +qed-. + +fact coafter_isfin2_fwd_aux: (âf1. @â¦0, f1⦠⡠0 â H_coafter_isfin2_fwd f1) â + âi2,f1. @â¦0, f1⦠⡠i2 â H_coafter_isfin2_fwd f1. +#H0 #i2 elim i2 -i2 /2 width=1 by/ -H0 +#i2 #IH #f1 #H1f1 #f2 #Hf2 #H2f1 #f #Hf +elim (at_inv_pxn ⦠H1f1) -H1f1 [ |*: // ] #g1 #Hg1 #H1 +elim (coafter_inv_nxx ⦠Hf ⦠H1) -Hf #g #Hg #H0 +lapply (IH ⦠Hg1 ⦠Hg) -i2 -Hg +/2 width=4 by istot_inv_next, isfin_push/ (**) (* full auto fails *) +qed-. + +lemma coafter_isfin2_fwd: âf1. H_coafter_isfin2_fwd f1. +#f1 #f2 #Hf2 #Hf1 cases (Hf1 0) +/3 width=7 by coafter_isfin2_fwd_aux, coafter_isfin2_fwd_O_aux/ +qed-. + lemma coafter_inv_sor: âf. ð â¦f⦠â âf2. ðâ¦f2⦠â âf1. f2 ~â f1 â¡ f â âfa,fb. fa â fb â¡ f â ââf1a,f1b. f2 ~â f1a â¡ fa & f2 ~â f1b â¡ fb & f1a â f1b â¡ f1. @isfin_ind @@ -606,4 +647,34 @@ lemma coafter_inv_sor: âf. ð â¦f⦠â âf2. ðâ¦f2⦠â âf1. f2 ~ elim (IH ⦠Hg2 ⦠H1f ⦠H2f) -f -Hg2 /3 width=11 by sor_np, sor_pn, sor_nn, ex3_2_intro, coafter_refl, coafter_push/ ] -qed-. +qed-. + +(* Properties with istot ****************************************************) + +lemma coafter_sor: âf. ð â¦f⦠â âf2. ðâ¦f2⦠â âf1. f2 ~â f1 â¡ f â âf1a,f1b. f1a â f1b â¡ f1 â + ââfa,fb. f2 ~â f1a â¡ fa & f2 ~â f1b â¡ fb & fa â fb â¡ f. +@isfin_ind +[ #f #Hf #f2 #Hf2 #f1 #Hf #f1a #f1b #Hf1 + lapply (coafter_fwd_isid2 ⦠Hf ??) -Hf // #H2f1 + elim (sor_inv_isid3 ⦠Hf1) -Hf1 // + /3 width=5 by coafter_isid_dx, sor_refl, ex3_2_intro/ +| #f #_ #IH #f2 #Hf2 #f1 #H1 #f1a #f1b #H2 + elim (coafter_inv_xxp ⦠H1) -H1 [1,3: * |*: // ] + [ #g2 #g1 #Hf #Hgf2 #Hgf1 + elim (sor_inv_xxp ⦠H2) -H2 [ |*: // ] #ga #gb #Hg1 + lapply (istot_inv_push ⦠Hf2 ⦠Hgf2) -Hf2 #Hg2 + elim (IH ⦠Hf ⦠Hg1) // -f1 -g1 -IH -Hg2 + /3 width=11 by coafter_refl, sor_pp, ex3_2_intro/ + | #g2 #Hf #Hgf2 + lapply (istot_inv_next ⦠Hf2 ⦠Hgf2) -Hf2 #Hg2 + elim (IH ⦠Hf ⦠H2) // -f1 -IH -Hg2 + /3 width=11 by coafter_next, sor_pp, ex3_2_intro/ + ] +| #f #_ #IH #f2 #Hf2 #f1 #H1 #f1a #f1b #H2 + elim (coafter_inv_xxn ⦠H1) -H1 [ |*: // ] #g2 #g1 #Hf #Hgf2 #Hgf1 + lapply (istot_inv_push ⦠Hf2 ⦠Hgf2) -Hf2 #Hg2 + elim (sor_inv_xxn ⦠H2) -H2 [1,3,4: * |*: // ] #ga #gb #Hg1 + elim (IH ⦠Hf ⦠Hg1) // -f1 -g1 -IH -Hg2 + /3 width=11 by coafter_refl, coafter_push, sor_np, sor_pn, sor_nn, ex3_2_intro/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isfin.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isfin.ma index 8a127b359..b73e34cea 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isfin.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isfin.ma @@ -74,5 +74,11 @@ qed. (* Inversion lemmas with tail ***********************************************) lemma isfin_inv_tl: âf. ð â¦â«±f⦠â ð â¦fâ¦. -#f elim (pn_split f) * /2 width=1 by isfin_next, isfin_push/ +#f elim (pn_split f) * /2 width=1 by isfin_next, isfin_push/ +qed-. + +(* Inversion lemmas with tls ********************************************************) + +lemma isfin_inv_tls: ân,f. ð â¦â«±*[n]f⦠â ð â¦fâ¦. +#n elim n -n /3 width=1 by isfin_inv_tl/ qed-. -- 2.39.5