X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground%2Frelocation%2Fgr_coafter.ma;h=2ec6991cf31f9e9f4ce6d2063ef847dbba8f8a6e;hb=8bbe582d87984526f40182c4409cbfd43108cb79;hp=8c730547fd1f8ca6123844ee1863e0dedb566c6d;hpb=55c768d7e45babb300b5010463ba3196a68f1bbe;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground/relocation/gr_coafter.ma b/matita/matita/contribs/lambdadelta/ground/relocation/gr_coafter.ma index 8c730547f..2ec6991cf 100644 --- a/matita/matita/contribs/lambdadelta/ground/relocation/gr_coafter.ma +++ b/matita/matita/contribs/lambdadelta/ground/relocation/gr_coafter.ma @@ -16,7 +16,7 @@ include "ground/notation/relations/rcoafter_3.ma". include "ground/xoa/ex_3_2.ma". include "ground/relocation/gr_tl.ma". -(* RELATIONAL CO-COMPOSITION FOR GENERIC RELOCATION MAPS ***********************************************************) +(* RELATIONAL CO-COMPOSITION FOR GENERIC RELOCATION MAPS ********************) (*** coafter *) coinductive gr_coafter: relation3 gr_map gr_map gr_map ≝ @@ -35,7 +35,7 @@ interpretation "relational co-composition (generic relocation maps)" 'RCoAfter f1 f2 f = (gr_coafter f1 f2 f). -(* Basic inversion lemmas ***************************************************) +(* Basic inversions *********************************************************) (*** coafter_inv_ppx *) lemma gr_coafter_inv_push_bi: @@ -82,7 +82,7 @@ lemma gr_coafter_inv_next_sn: ] qed-. -(* Advanced inversion lemmas ************************************************) +(* Advanced inversions ******************************************************) (*** coafter_inv_ppp *) lemma gr_coafter_inv_push_bi_push: @@ -213,26 +213,26 @@ elim (gr_map_split_tl g2) #H2 ] qed-. -(* Inversion lemmas with tail ***********************************************) +(* Inversions with gr_tl ****************************************************) (*** coafter_inv_tl1 *) lemma gr_coafter_inv_tl_dx: - ∀g2,g1,g. g2 ~⊚ ⫱g1 ≘ g → - ∃∃f. ⫯g2 ~⊚ g1 ≘ f & ⫱f = g. + ∀g2,g1,g. g2 ~⊚ ⫰g1 ≘ g → + ∃∃f. ⫯g2 ~⊚ g1 ≘ f & ⫰f = g. #g2 #g1 #g elim (gr_map_split_tl g1) #H1 #H2 [ /3 width=7 by gr_coafter_refl, ex2_intro/ -| @(ex2_intro … (↑g)) /2 width=7 by gr_coafter_push/ (**) (* full auto fails *) +| @(ex2_intro … (↑g)) /2 width=7 by gr_coafter_push/ (* * full auto fails *) ] qed-. (*** coafter_inv_tl0 *) lemma gr_coafter_inv_tl: - ∀g2,g1,g. g2 ~⊚ g1 ≘ ⫱g → - ∃∃f1. ⫯g2 ~⊚ f1 ≘ g & ⫱f1 = g1. + ∀g2,g1,g. g2 ~⊚ g1 ≘ ⫰g → + ∃∃f1. ⫯g2 ~⊚ f1 ≘ g & ⫰f1 = g1. #g2 #g1 #g elim (gr_map_split_tl g) #H1 #H2 [ /3 width=7 by gr_coafter_refl, ex2_intro/ -| @(ex2_intro … (↑g1)) /2 width=7 by gr_coafter_push/ (**) (* full auto fails *) +| @(ex2_intro … (↑g1)) /2 width=7 by gr_coafter_push/ (* * full auto fails *) ] qed-.