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=2fc90abd113e3dbecbac60066707a03fc983a0b0;hpb=045c74915022181e288d9a950cc485437b08d002;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 2fc90abd1..fb59fcb78 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sle.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sle.ma @@ -13,6 +13,7 @@ (**************************************************************************) include "ground_2/relocation/rtmap_isid.ma". +include "ground_2/relocation/rtmap_isdiv.ma". (* RELOCATION MAP ***********************************************************) @@ -23,15 +24,30 @@ coinductive sle: relation rtmap ≝ . interpretation "inclusion (rtmap)" - 'subseteq t1 t2 = (sle t1 t2). + 'subseteq f1 f2 = (sle f1 f2). (* Basic properties *********************************************************) +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 → @@ -67,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-. @@ -79,6 +95,13 @@ lemma sle_inv_px: ∀g1,g2. g1 ⊆ g2 → ∀f1. ↑f1 = g1 → /3 width=3 by ex2_intro, or_introl, or_intror/ qed-. +lemma sle_inv_xn: ∀g1,g2. g1 ⊆ g2 → ∀f2. ⫯f2 = g2 → + (∃∃f1. f1 ⊆ f2 & ↑f1 = g1) ∨ ∃∃f1. f1 ⊆ f2 & ⫯f1 = g1. +#g1 #g2 elim (pn_split g1) * #f1 #H1 #H #f2 #H2 +[ lapply (sle_inv_pn … H … H1 H2) | lapply (sle_inv_nn … H … H1 H2) ] -H -H2 +/3 width=3 by ex2_intro, or_introl, or_intror/ +qed-. + (* Main properties **********************************************************) corec theorem sle_trans: Transitive … sle. @@ -100,6 +123,10 @@ lemma sle_px_tl: ∀g1,g2. g1 ⊆ g2 → ∀f1. ↑f1 = g1 → f1 ⊆ ⫱g2. #g1 #g2 #H #f1 #H1 elim (sle_inv_px … H … H1) -H -H1 * // qed. +lemma sle_xn_tl: ∀g1,g2. g1 ⊆ g2 → ∀f2. ⫯f2 = g2 → ⫱g1 ⊆ f2. +#g1 #g2 #H #f2 #H2 elim (sle_inv_xn … H … H2) -H -H2 * // +qed. + lemma sle_tl: ∀f1,f2. f1 ⊆ f2 → ⫱f1 ⊆ ⫱f2. #f1 elim (pn_split f1) * #g1 #H1 #f2 #H [ lapply (sle_px_tl … H … H1) -H // @@ -119,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. @@ -136,3 +169,21 @@ corec lemma sle_inv_isid_dx: ∀f1,f2. f1 ⊆ f2 → 𝐈⦃f2⦄ → 𝐈⦃f1 lapply (isid_inv_push … H ??) -H /3 width=3 by isid_push/ qed-. + +(* Properties with isdiv ****************************************************) + +corec lemma sle_isdiv_dx: ∀f2. 𝛀⦃f2⦄ → ∀f1. f1 ⊆ f2. +#f2 * -f2 +#f2 #g2 #Hf2 #H2 #f1 cases (pn_split f1) * +/3 width=5 by sle_weak, sle_next/ +qed. + +(* Inversion lemmas with isdiv **********************************************) + +corec lemma sle_inv_isdiv_sn: ∀f1,f2. f1 ⊆ f2 → 𝛀⦃f1⦄ → 𝛀⦃f2⦄. +#f1 #f2 * -f1 -f2 +#f1 #f2 #g1 #g2 #Hf * * #H +[1,3: elim (isdiv_inv_push … H) // ] +lapply (isdiv_inv_next … H ??) -H +/3 width=3 by isdiv_next/ +qed-.