X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground%2Frelocation%2Fgr_sle.ma;h=74c15789f00bbc773189e17d68b58a735bfe6bf2;hb=8bbe582d87984526f40182c4409cbfd43108cb79;hp=6ad077f23c028333f940746355c058313cb55c21;hpb=55c768d7e45babb300b5010463ba3196a68f1bbe;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground/relocation/gr_sle.ma b/matita/matita/contribs/lambdadelta/ground/relocation/gr_sle.ma index 6ad077f23..74c15789f 100644 --- a/matita/matita/contribs/lambdadelta/ground/relocation/gr_sle.ma +++ b/matita/matita/contribs/lambdadelta/ground/relocation/gr_sle.ma @@ -14,7 +14,7 @@ include "ground/relocation/gr_tl.ma". -(* INCLUSION FOR GENERIC RELOCATION MAPS ***********************************************************) +(* INCLUSION FOR GENERIC RELOCATION MAPS ************************************) (*** sle *) coinductive gr_sle: relation gr_map ≝ @@ -33,7 +33,7 @@ interpretation "inclusion (generic relocation maps)" 'subseteq f1 f2 = (gr_sle f1 f2). -(* Basic properties *********************************************************) +(* Basic constructions ******************************************************) (*** sle_refl *) corec lemma gr_sle_refl: @@ -42,7 +42,7 @@ corec lemma gr_sle_refl: [ @(gr_sle_push … H H) | @(gr_sle_next … H H) ] -H // qed. -(* Basic inversion lemmas ***************************************************) +(* Basic inversions *********************************************************) (*** sle_inv_xp *) lemma gr_sle_inv_push_dx: @@ -76,7 +76,7 @@ lemma gr_sle_inv_push_next: ] qed-. -(* Advanced inversion lemmas ************************************************) +(* Advanced inversions ******************************************************) (*** sle_inv_pp *) lemma gr_sle_inv_push_bi: @@ -120,43 +120,43 @@ elim (gr_map_split_tl g1) #H1 #H #f2 #H2 /3 width=3 by ex2_intro, or_introl, or_intror/ qed-. -(* Properties with tail *****************************************************) +(* Constructions with gr_tl *************************************************) (*** sle_px_tl *) lemma gr_sle_push_sn_tl: - ∀g1,g2. g1 ⊆ g2 → ∀f1. ⫯f1 = g1 → f1 ⊆ ⫱g2. + ∀g1,g2. g1 ⊆ g2 → ∀f1. ⫯f1 = g1 → f1 ⊆ ⫰g2. #g1 #g2 #H #f1 #H1 elim (gr_sle_inv_push_sn … H … H1) -H -H1 * // qed. (*** sle_xn_tl *) lemma gr_sle_next_dx_tl: - ∀g1,g2. g1 ⊆ g2 → ∀f2. ↑f2 = g2 → ⫱g1 ⊆ f2. + ∀g1,g2. g1 ⊆ g2 → ∀f2. ↑f2 = g2 → ⫰g1 ⊆ f2. #g1 #g2 #H #f2 #H2 elim (gr_sle_inv_next_dx … H … H2) -H -H2 * // qed. (*** sle_tl *) lemma gr_sle_tl: - ∀f1,f2. f1 ⊆ f2 → ⫱f1 ⊆ ⫱f2. + ∀f1,f2. f1 ⊆ f2 → ⫰f1 ⊆ ⫰f2. #f1 elim (gr_map_split_tl f1) #H1 #f2 #H [ lapply (gr_sle_push_sn_tl … H … H1) -H // | elim (gr_sle_inv_next_sn … H … H1) -H // ] qed. -(* Inversion lemmas with tail ***********************************************) +(* Inversions with gr_tl ****************************************************) (*** sle_inv_tl_sn *) lemma gr_sle_inv_tl_sn: - ∀f1,f2. ⫱f1 ⊆ f2 → f1 ⊆ ↑f2. + ∀f1,f2. ⫰f1 ⊆ f2 → f1 ⊆ ↑f2. #f1 elim (gr_map_split_tl f1) #H #f2 #Hf12 /2 width=5 by gr_sle_next, gr_sle_weak/ qed-. (*** sle_inv_tl_dx *) lemma gr_sle_inv_tl_dx: - ∀f1,f2. f1 ⊆ ⫱f2 → ⫯f1 ⊆ f2. + ∀f1,f2. f1 ⊆ ⫰f2 → ⫯f1 ⊆ f2. #f1 #f2 elim (gr_map_split_tl f2) #H #Hf12 /2 width=5 by gr_sle_push, gr_sle_weak/ qed-.