X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground%2Frelocation%2Fgr_sor.ma;h=4f2e759e03f5b4b9bf646f9469c800e619ede9ab;hb=8bbe582d87984526f40182c4409cbfd43108cb79;hp=7152502c2016465a176dac7b73194a6d0d634f45;hpb=55c768d7e45babb300b5010463ba3196a68f1bbe;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground/relocation/gr_sor.ma b/matita/matita/contribs/lambdadelta/ground/relocation/gr_sor.ma index 7152502c2..4f2e759e0 100644 --- a/matita/matita/contribs/lambdadelta/ground/relocation/gr_sor.ma +++ b/matita/matita/contribs/lambdadelta/ground/relocation/gr_sor.ma @@ -17,7 +17,7 @@ include "ground/xoa/or_3.ma". include "ground/xoa/ex_3_2.ma". include "ground/relocation/gr_tl.ma". -(* RELATIONAL UNION FOR GENERIC RELOCATION MAPS ***********************************************************) +(* RELATIONAL UNION FOR GENERIC RELOCATION MAPS *****************************) (*** sor *) coinductive gr_sor: relation3 gr_map gr_map gr_map ≝ @@ -39,7 +39,7 @@ interpretation "relational union (generic relocation maps)" 'RUnion f1 f2 f = (gr_sor f1 f2 f). -(* Basic properties *********************************************************) +(* Basic constructions ******************************************************) (*** sor_idem *) corec lemma gr_sor_idem: @@ -58,7 +58,7 @@ corec lemma gr_sor_comm: [ @gr_sor_push_bi | @gr_sor_push_next | @gr_sor_next_push | @gr_sor_next_bi ] /2 width=7 by/ qed-. -(* Basic inversion lemmas ***************************************************) +(* Basic inversions *********************************************************) (*** sor_inv_ppx *) lemma gr_sor_inv_push_bi: @@ -112,7 +112,7 @@ try elim (eq_inv_gr_push_next … Hx2) try elim (eq_inv_gr_next_push … Hx2) /2 width=3 by ex2_intro/ qed-. -(* Advanced inversion lemmas ************************************************) +(* Advanced inversions ******************************************************) (*** sor_inv_ppn *) lemma gr_sor_inv_push_bi_next: @@ -277,11 +277,11 @@ elim (gr_map_split_tl g1) #H1 ] qed-. -(* Properties with tail *****************************************************) +(* Constructions with gr_tl *************************************************) (*** sor_tl *) lemma gr_sor_tl: - ∀f1,f2,f. f1 ⋓ f2 ≘ f → ⫱f1 ⋓ ⫱f2 ≘ ⫱f. + ∀f1,f2,f. f1 ⋓ f2 ≘ f → ⫰f1 ⋓ ⫰f2 ≘ ⫰f. #f1 cases (gr_map_split_tl f1) #H1 #f2 cases (gr_map_split_tl f2) #H2 #f #Hf @@ -295,8 +295,8 @@ qed. (*** sor_xxn_tl *) lemma gr_sor_next_tl: ∀g1,g2,g. g1 ⋓ g2 ≘ g → ∀f. ↑f = g → - (∃∃f1,f2. f1 ⋓ f2 ≘ f & ↑f1 = g1 & ⫱g2 = f2) ∨ - (∃∃f1,f2. f1 ⋓ f2 ≘ f & ⫱g1 = f1 & ↑f2 = g2). + (∃∃f1,f2. f1 ⋓ f2 ≘ f & ↑f1 = g1 & ⫰g2 = f2) ∨ + (∃∃f1,f2. f1 ⋓ f2 ≘ f & ⫰g1 = f1 & ↑f2 = g2). #g1 #g2 #g #H #f #H0 elim (gr_sor_inv_next … H … H0) -H -H0 * /3 width=5 by ex3_2_intro, or_introl, or_intror/ qed-. @@ -304,7 +304,7 @@ qed-. (*** sor_xnx_tl *) lemma gr_sor_next_dx_tl: ∀g1,g2,g. g1 ⋓ g2 ≘ g → ∀f2. ↑f2 = g2 → - ∃∃f1,f. f1 ⋓ f2 ≘ f & ⫱g1 = f1 & ↑f = g. + ∃∃f1,f. f1 ⋓ f2 ≘ f & ⫰g1 = f1 & ↑f = g. #g1 elim (gr_map_split_tl g1) #H1 #g2 #g #H #f2 #H2 [ elim (gr_sor_inv_push_next … H … H1 H2) | elim (gr_sor_inv_next_bi … H … H1 H2) @@ -315,7 +315,7 @@ qed-. (*** sor_nxx_tl *) lemma gr_sor_next_sn_tl: ∀g1,g2,g. g1 ⋓ g2 ≘ g → ∀f1. ↑f1 = g1 → - ∃∃f2,f. f1 ⋓ f2 ≘ f & ⫱g2 = f2 & ↑f = g. + ∃∃f2,f. f1 ⋓ f2 ≘ f & ⫰g2 = f2 & ↑f = g. #g1 #g2 elim (gr_map_split_tl g2) #H2 #g #H #f1 #H1 [ elim (gr_sor_inv_next_push … H … H1 H2) | elim (gr_sor_inv_next_bi … H … H1 H2) @@ -323,18 +323,18 @@ lemma gr_sor_next_sn_tl: /3 width=5 by ex3_2_intro/ qed-. -(* Inversion lemmas with tail ***********************************************) +(* Inversions with gr_tl ****************************************************) (*** sor_inv_tl_sn *) lemma gr_sor_inv_tl_sn: - ∀f1,f2,f. ⫱f1 ⋓ f2 ≘ f → f1 ⋓ ↑f2 ≘ ↑f. + ∀f1,f2,f. ⫰f1 ⋓ f2 ≘ f → f1 ⋓ ↑f2 ≘ ↑f. #f1 #f2 #f elim (gr_map_split_tl f1) /2 width=7 by gr_sor_push_next, gr_sor_next_bi/ qed-. (*** sor_inv_tl_dx *) lemma gr_sor_inv_tl_dx: - ∀f1,f2,f. f1 ⋓ ⫱f2 ≘ f → ↑f1 ⋓ f2 ≘ ↑f. + ∀f1,f2,f. f1 ⋓ ⫰f2 ≘ f → ↑f1 ⋓ f2 ≘ ↑f. #f1 #f2 #f elim (gr_map_split_tl f2) /2 width=7 by gr_sor_next_push, gr_sor_next_bi/ qed-.