X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Frelocation%2Frtmap_sle.ma;h=fb59fcb786e84c360886d3642fdf2843e9ae6660;hb=b1868c5a258a6bf7fc983d63f3c417f00185e7b6;hp=cebc3aa4a463d03f786933939c3d890999c25f43;hpb=66962864d3703b8f3b44e95d32c03ed50ceee6f1;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sle.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sle.ma index cebc3aa4a..fb59fcb78 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sle.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sle.ma @@ -24,16 +24,30 @@ coinductive sle: relation rtmap ≝ . interpretation "inclusion (rtmap)" - 'subseteq t1 t2 = (sle t1 t2). + 'subseteq f1 f2 = (sle f1 f2). (* Basic properties *********************************************************) -corec lemma sle_refl: ∀f1,f2. f1 ≗ f2 → f1 ⊆ f2. -#f1 #f2 * -f1 -f2 -#f1 #f2 #g1 #g2 #H12 #H1 #H2 -[ @(sle_push … H1 H2) | @(sle_next … H1 H2) ] -H1 -H2 /2 width=1 by/ +axiom sle_eq_repl_back1: ∀f2. eq_repl_back … (λf1. f1 ⊆ f2). + +lemma sle_eq_repl_fwd1: ∀f2. eq_repl_fwd … (λf1. f1 ⊆ f2). +#f2 @eq_repl_sym /2 width=3 by sle_eq_repl_back1/ +qed-. + +axiom sle_eq_repl_back2: ∀f1. eq_repl_back … (λf2. f1 ⊆ f2). + +lemma sle_eq_repl_fwd2: ∀f1. eq_repl_fwd … (λf2. f1 ⊆ f2). +#f1 @eq_repl_sym /2 width=3 by sle_eq_repl_back2/ +qed-. + +corec lemma sle_refl: ∀f. f ⊆ f. +#f cases (pn_split f) * #g #H +[ @(sle_push … H H) | @(sle_next … H H) ] -H // qed. +lemma sle_refl_eq: ∀f1,f2. f1 ≗ f2 → f1 ⊆ f2. +/2 width=3 by sle_eq_repl_back2/ qed. + (* Basic inversion lemmas ***************************************************) lemma sle_inv_xp: ∀g1,g2. g1 ⊆ g2 → ∀f2. ↑f2 = g2 → @@ -69,7 +83,7 @@ lemma sle_inv_pp: ∀g1,g2. g1 ⊆ g2 → ∀f1,f2. ↑f1 = g1 → ↑f2 = g2 #x1 #H #Hx1 destruct lapply (injective_push … Hx1) -Hx1 // qed-. -lemma sle_inv_nn: ∀g1,g2. g1 ⊆ g2 → ∀f1,f2. ⫯f1 = g1 → ⫯f2 = g2 → f1 ⊆ f2. +lemma sle_inv_nn: ∀g1,g2. g1 ⊆ g2 → ∀f1,f2. ⫯f1 = g1 → ⫯f2 = g2 → f1 ⊆ f2. #g1 #g2 #H #f1 #f2 #H1 #H2 elim (sle_inv_nx … H … H1) -g1 #x2 #H #Hx2 destruct lapply (injective_next … Hx2) -Hx2 // qed-. @@ -132,6 +146,12 @@ lemma sle_inv_tl_dx: ∀f1,f2. f1 ⊆ ⫱f2 → ↑f1 ⊆ f2. /2 width=5 by sle_push, sle_weak/ qed-. +(* Properties with iteraded tail ********************************************) + +lemma sle_tls: ∀f1,f2. f1 ⊆ f2 → ∀i. ⫱*[i] f1 ⊆ ⫱*[i] f2. +#f1 #f2 #Hf12 #i elim i -i /2 width=5 by sle_tl/ +qed. + (* Properties with isid *****************************************************) corec lemma sle_isid_sn: ∀f1. 𝐈⦃f1⦄ → ∀f2. f1 ⊆ f2.