X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Frelocation%2Frtmap_sor.ma;h=9191eb3212b6aa67185a8f444a3c7fc2fba45af8;hb=bd53c4e895203eb049e75434f638f26b5a161a2b;hp=b93214beadafbf4198d5d64d27bda5c9009ede38;hpb=ff1cd6f29b3aaef01e4674544d399f44949c5738;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sor.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sor.ma index b93214bea..9191eb321 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sor.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sor.ma @@ -12,15 +12,16 @@ (* *) (**************************************************************************) +include "ground_2/xoa/ex_4_2.ma". include "ground_2/notation/relations/runion_3.ma". include "ground_2/relocation/rtmap_isfin.ma". include "ground_2/relocation/rtmap_sle.ma". coinductive sor: relation3 rtmap rtmap rtmap ≝ -| sor_pp: ∀f1,f2,f,g1,g2,g. sor f1 f2 f → ↑f1 = g1 → ↑f2 = g2 → ↑f = g → sor g1 g2 g -| sor_np: ∀f1,f2,f,g1,g2,g. sor f1 f2 f → ⫯f1 = g1 → ↑f2 = g2 → ⫯f = g → sor g1 g2 g -| sor_pn: ∀f1,f2,f,g1,g2,g. sor f1 f2 f → ↑f1 = g1 → ⫯f2 = g2 → ⫯f = g → sor g1 g2 g -| sor_nn: ∀f1,f2,f,g1,g2,g. sor f1 f2 f → ⫯f1 = g1 → ⫯f2 = g2 → ⫯f = g → sor g1 g2 g +| sor_pp: ∀f1,f2,f,g1,g2,g. sor f1 f2 f → ⫯f1 = g1 → ⫯f2 = g2 → ⫯f = g → sor g1 g2 g +| sor_np: ∀f1,f2,f,g1,g2,g. sor f1 f2 f → ↑f1 = g1 → ⫯f2 = g2 → ↑f = g → sor g1 g2 g +| sor_pn: ∀f1,f2,f,g1,g2,g. sor f1 f2 f → ⫯f1 = g1 → ↑f2 = g2 → ↑f = g → sor g1 g2 g +| sor_nn: ∀f1,f2,f,g1,g2,g. sor f1 f2 f → ↑f1 = g1 → ↑f2 = g2 → ↑f = g → sor g1 g2 g . interpretation "union (rtmap)" @@ -28,8 +29,8 @@ interpretation "union (rtmap)" (* Basic inversion lemmas ***************************************************) -lemma sor_inv_ppx: ∀g1,g2,g. g1 ⋓ g2 ≡ g → ∀f1,f2. ↑f1 = g1 → ↑f2 = g2 → - ∃∃f. f1 ⋓ f2 ≡ f & ↑f = g. +lemma sor_inv_ppx: ∀g1,g2,g. g1 ⋓ g2 ≘ g → ∀f1,f2. ⫯f1 = g1 → ⫯f2 = g2 → + ∃∃f. f1 ⋓ f2 ≘ f & ⫯f = g. #g1 #g2 #g * -g1 -g2 -g #f1 #f2 #f #g1 #g2 #g #Hf #H1 #H2 #H0 #x1 #x2 #Hx1 #Hx2 destruct try (>(injective_push … Hx1) -x1) try (>(injective_next … Hx1) -x1) @@ -39,8 +40,8 @@ try elim (discr_push_next … Hx2) try elim (discr_next_push … Hx2) /2 width=3 by ex2_intro/ qed-. -lemma sor_inv_npx: ∀g1,g2,g. g1 ⋓ g2 ≡ g → ∀f1,f2. ⫯f1 = g1 → ↑f2 = g2 → - ∃∃f. f1 ⋓ f2 ≡ f & ⫯f = g. +lemma sor_inv_npx: ∀g1,g2,g. g1 ⋓ g2 ≘ g → ∀f1,f2. ↑f1 = g1 → ⫯f2 = g2 → + ∃∃f. f1 ⋓ f2 ≘ f & ↑f = g. #g1 #g2 #g * -g1 -g2 -g #f1 #f2 #f #g1 #g2 #g #Hf #H1 #H2 #H0 #x1 #x2 #Hx1 #Hx2 destruct try (>(injective_push … Hx1) -x1) try (>(injective_next … Hx1) -x1) @@ -50,8 +51,8 @@ try elim (discr_push_next … Hx2) try elim (discr_next_push … Hx2) /2 width=3 by ex2_intro/ qed-. -lemma sor_inv_pnx: ∀g1,g2,g. g1 ⋓ g2 ≡ g → ∀f1,f2. ↑f1 = g1 → ⫯f2 = g2 → - ∃∃f. f1 ⋓ f2 ≡ f & ⫯f = g. +lemma sor_inv_pnx: ∀g1,g2,g. g1 ⋓ g2 ≘ g → ∀f1,f2. ⫯f1 = g1 → ↑f2 = g2 → + ∃∃f. f1 ⋓ f2 ≘ f & ↑f = g. #g1 #g2 #g * -g1 -g2 -g #f1 #f2 #f #g1 #g2 #g #Hf #H1 #H2 #H0 #x1 #x2 #Hx1 #Hx2 destruct try (>(injective_push … Hx1) -x1) try (>(injective_next … Hx1) -x1) @@ -61,8 +62,8 @@ try elim (discr_push_next … Hx2) try elim (discr_next_push … Hx2) /2 width=3 by ex2_intro/ qed-. -lemma sor_inv_nnx: ∀g1,g2,g. g1 ⋓ g2 ≡ g → ∀f1,f2. ⫯f1 = g1 → ⫯f2 = g2 → - ∃∃f. f1 ⋓ f2 ≡ f & ⫯f = g. +lemma sor_inv_nnx: ∀g1,g2,g. g1 ⋓ g2 ≘ g → ∀f1,f2. ↑f1 = g1 → ↑f2 = g2 → + ∃∃f. f1 ⋓ f2 ≘ f & ↑f = g. #g1 #g2 #g * -g1 -g2 -g #f1 #f2 #f #g1 #g2 #g #Hf #H1 #H2 #H0 #x1 #x2 #Hx1 #Hx2 destruct try (>(injective_push … Hx1) -x1) try (>(injective_next … Hx1) -x1) @@ -72,9 +73,144 @@ try elim (discr_push_next … Hx2) try elim (discr_next_push … Hx2) /2 width=3 by ex2_intro/ qed-. +(* Advanced inversion lemmas ************************************************) + +lemma sor_inv_ppn: ∀g1,g2,g. g1 ⋓ g2 ≘ g → + ∀f1,f2,f. ⫯f1 = g1 → ⫯f2 = g2 → ↑f = g → ⊥. +#g1 #g2 #g #H #f1 #f2 #f #H1 #H2 #H0 +elim (sor_inv_ppx … H … H1 H2) -g1 -g2 #x #_ #H destruct +/2 width=3 by discr_push_next/ +qed-. + +lemma sor_inv_nxp: ∀g1,g2,g. g1 ⋓ g2 ≘ g → + ∀f1,f. ↑f1 = g1 → ⫯f = g → ⊥. +#g1 #g2 #g #H #f1 #f #H1 #H0 +elim (pn_split g2) * #f2 #H2 +[ elim (sor_inv_npx … H … H1 H2) +| elim (sor_inv_nnx … H … H1 H2) +] -g1 -g2 #x #_ #H destruct +/2 width=3 by discr_next_push/ +qed-. + +lemma sor_inv_xnp: ∀g1,g2,g. g1 ⋓ g2 ≘ g → + ∀f2,f. ↑f2 = g2 → ⫯f = g → ⊥. +#g1 #g2 #g #H #f2 #f #H2 #H0 +elim (pn_split g1) * #f1 #H1 +[ elim (sor_inv_pnx … H … H1 H2) +| elim (sor_inv_nnx … H … H1 H2) +] -g1 -g2 #x #_ #H destruct +/2 width=3 by discr_next_push/ +qed-. + +lemma sor_inv_ppp: ∀g1,g2,g. g1 ⋓ g2 ≘ g → + ∀f1,f2,f. ⫯f1 = g1 → ⫯f2 = g2 → ⫯f = g → f1 ⋓ f2 ≘ f. +#g1 #g2 #g #H #f1 #f2 #f #H1 #H2 #H0 +elim (sor_inv_ppx … H … H1 H2) -g1 -g2 #x #Hx #H destruct +<(injective_push … H) -f // +qed-. + +lemma sor_inv_npn: ∀g1,g2,g. g1 ⋓ g2 ≘ g → + ∀f1,f2,f. ↑f1 = g1 → ⫯f2 = g2 → ↑f = g → f1 ⋓ f2 ≘ f. +#g1 #g2 #g #H #f1 #f2 #f #H1 #H2 #H0 +elim (sor_inv_npx … H … H1 H2) -g1 -g2 #x #Hx #H destruct +<(injective_next … H) -f // +qed-. + +lemma sor_inv_pnn: ∀g1,g2,g. g1 ⋓ g2 ≘ g → + ∀f1,f2,f. ⫯f1 = g1 → ↑f2 = g2 → ↑f = g → f1 ⋓ f2 ≘ f. +#g1 #g2 #g #H #f1 #f2 #f #H1 #H2 #H0 +elim (sor_inv_pnx … H … H1 H2) -g1 -g2 #x #Hx #H destruct +<(injective_next … H) -f // +qed-. + +lemma sor_inv_nnn: ∀g1,g2,g. g1 ⋓ g2 ≘ g → + ∀f1,f2,f. ↑f1 = g1 → ↑f2 = g2 → ↑f = g → f1 ⋓ f2 ≘ f. +#g1 #g2 #g #H #f1 #f2 #f #H1 #H2 #H0 +elim (sor_inv_nnx … H … H1 H2) -g1 -g2 #x #Hx #H destruct +<(injective_next … H) -f // +qed-. + +lemma sor_inv_pxp: ∀g1,g2,g. g1 ⋓ g2 ≘ g → + ∀f1,f. ⫯f1 = g1 → ⫯f = g → + ∃∃f2. f1 ⋓ f2 ≘ f & ⫯f2 = g2. +#g1 #g2 #g #H #f1 #f #H1 #H0 +elim (pn_split g2) * #f2 #H2 +[ /3 width=7 by sor_inv_ppp, ex2_intro/ +| elim (sor_inv_xnp … H … H2 H0) +] +qed-. + +lemma sor_inv_xpp: ∀g1,g2,g. g1 ⋓ g2 ≘ g → + ∀f2,f. ⫯f2 = g2 → ⫯f = g → + ∃∃f1. f1 ⋓ f2 ≘ f & ⫯f1 = g1. +#g1 #g2 #g #H #f2 #f #H2 #H0 +elim (pn_split g1) * #f1 #H1 +[ /3 width=7 by sor_inv_ppp, ex2_intro/ +| elim (sor_inv_nxp … H … H1 H0) +] +qed-. + +lemma sor_inv_pxn: ∀g1,g2,g. g1 ⋓ g2 ≘ g → + ∀f1,f. ⫯f1 = g1 → ↑f = g → + ∃∃f2. f1 ⋓ f2 ≘ f & ↑f2 = g2. +#g1 #g2 #g #H #f1 #f #H1 #H0 +elim (pn_split g2) * #f2 #H2 +[ elim (sor_inv_ppn … H … H1 H2 H0) +| /3 width=7 by sor_inv_pnn, ex2_intro/ +] +qed-. + +lemma sor_inv_xpn: ∀g1,g2,g. g1 ⋓ g2 ≘ g → + ∀f2,f. ⫯f2 = g2 → ↑f = g → + ∃∃f1. f1 ⋓ f2 ≘ f & ↑f1 = g1. +#g1 #g2 #g #H #f2 #f #H2 #H0 +elim (pn_split g1) * #f1 #H1 +[ elim (sor_inv_ppn … H … H1 H2 H0) +| /3 width=7 by sor_inv_npn, ex2_intro/ +] +qed-. + +lemma sor_inv_xxp: ∀g1,g2,g. g1 ⋓ g2 ≘ g → ∀f. ⫯f = g → + ∃∃f1,f2. f1 ⋓ f2 ≘ f & ⫯f1 = g1 & ⫯f2 = g2. +#g1 #g2 #g #H #f #H0 +elim (pn_split g1) * #f1 #H1 +[ elim (sor_inv_pxp … H … H1 H0) -g /2 width=5 by ex3_2_intro/ +| elim (sor_inv_nxp … H … H1 H0) +] +qed-. + +lemma sor_inv_nxn: ∀g1,g2,g. g1 ⋓ g2 ≘ g → + ∀f1,f. ↑f1 = g1 → ↑f = g → + (∃∃f2. f1 ⋓ f2 ≘ f & ⫯f2 = g2) ∨ + ∃∃f2. f1 ⋓ f2 ≘ f & ↑f2 = g2. +#g1 #g2 elim (pn_split g2) * +/4 width=7 by sor_inv_npn, sor_inv_nnn, ex2_intro, or_intror, or_introl/ +qed-. + +lemma sor_inv_xnn: ∀g1,g2,g. g1 ⋓ g2 ≘ g → + ∀f2,f. ↑f2 = g2 → ↑f = g → + (∃∃f1. f1 ⋓ f2 ≘ f & ⫯f1 = g1) ∨ + ∃∃f1. f1 ⋓ f2 ≘ f & ↑f1 = g1. +#g1 elim (pn_split g1) * +/4 width=7 by sor_inv_pnn, sor_inv_nnn, ex2_intro, or_intror, or_introl/ +qed-. + +lemma sor_inv_xxn: ∀g1,g2,g. g1 ⋓ g2 ≘ g → ∀f. ↑f = g → + ∨∨ ∃∃f1,f2. f1 ⋓ f2 ≘ f & ↑f1 = g1 & ⫯f2 = g2 + | ∃∃f1,f2. f1 ⋓ f2 ≘ f & ⫯f1 = g1 & ↑f2 = g2 + | ∃∃f1,f2. f1 ⋓ f2 ≘ f & ↑f1 = g1 & ↑f2 = g2. +#g1 #g2 #g #H #f #H0 +elim (pn_split g1) * #f1 #H1 +[ elim (sor_inv_pxn … H … H1 H0) -g + /3 width=5 by or3_intro1, ex3_2_intro/ +| elim (sor_inv_nxn … H … H1 H0) -g * + /3 width=5 by or3_intro0, or3_intro2, ex3_2_intro/ +] +qed-. + (* Main inversion lemmas ****************************************************) -corec theorem sor_mono: ∀f1,f2,x,y. f1 ⋓ f2 ≡ x → f1 ⋓ f2 ≡ y → x ≗ y. +corec theorem sor_mono: ∀f1,f2,x,y. f1 ⋓ f2 ≘ x → f1 ⋓ f2 ≘ y → x ≡ y. #f1 #f2 #x #y * -f1 -f2 -x #f1 #f2 #f #g1 #g2 #g #Hf #H1 #H2 #H0 #H [ cases (sor_inv_ppx … H … H1 H2) @@ -87,92 +223,330 @@ qed-. (* Basic properties *********************************************************) -corec lemma sor_eq_repl_back1: ∀f2,f. eq_repl_back … (λf1. f1 ⋓ f2 ≡ f). +corec lemma sor_eq_repl_back1: ∀f2,f. eq_repl_back … (λf1. f1 ⋓ f2 ≘ f). #f2 #f #f1 * -f1 -f2 -f #f1 #f2 #f #g1 #g2 #g #Hf #H1 #H2 #H0 #x #Hx try cases (eq_inv_px … Hx … H1) try cases (eq_inv_nx … Hx … H1) -g1 /3 width=7 by sor_pp, sor_np, sor_pn, sor_nn/ qed-. -lemma sor_eq_repl_fwd1: ∀f2,f. eq_repl_fwd … (λf1. f1 ⋓ f2 ≡ f). +lemma sor_eq_repl_fwd1: ∀f2,f. eq_repl_fwd … (λf1. f1 ⋓ f2 ≘ f). #f2 #f @eq_repl_sym /2 width=3 by sor_eq_repl_back1/ qed-. -corec lemma sor_eq_repl_back2: ∀f1,f. eq_repl_back … (λf2. f1 ⋓ f2 ≡ f). +corec lemma sor_eq_repl_back2: ∀f1,f. eq_repl_back … (λf2. f1 ⋓ f2 ≘ f). #f1 #f #f2 * -f1 -f2 -f #f1 #f2 #f #g1 #g2 #g #Hf #H #H2 #H0 #x #Hx try cases (eq_inv_px … Hx … H2) try cases (eq_inv_nx … Hx … H2) -g2 /3 width=7 by sor_pp, sor_np, sor_pn, sor_nn/ qed-. -lemma sor_eq_repl_fwd2: ∀f1,f. eq_repl_fwd … (λf2. f1 ⋓ f2 ≡ f). +lemma sor_eq_repl_fwd2: ∀f1,f. eq_repl_fwd … (λf2. f1 ⋓ f2 ≘ f). #f1 #f @eq_repl_sym /2 width=3 by sor_eq_repl_back2/ qed-. -corec lemma sor_eq_repl_back3: ∀f1,f2. eq_repl_back … (λf. f1 ⋓ f2 ≡ f). +corec lemma sor_eq_repl_back3: ∀f1,f2. eq_repl_back … (λf. f1 ⋓ f2 ≘ f). #f1 #f2 #f * -f1 -f2 -f #f1 #f2 #f #g1 #g2 #g #Hf #H #H2 #H0 #x #Hx try cases (eq_inv_px … Hx … H0) try cases (eq_inv_nx … Hx … H0) -g /3 width=7 by sor_pp, sor_np, sor_pn, sor_nn/ qed-. -lemma sor_eq_repl_fwd3: ∀f1,f2. eq_repl_fwd … (λf. f1 ⋓ f2 ≡ f). +lemma sor_eq_repl_fwd3: ∀f1,f2. eq_repl_fwd … (λf. f1 ⋓ f2 ≘ f). #f1 #f2 @eq_repl_sym /2 width=3 by sor_eq_repl_back3/ qed-. -corec lemma sor_refl: ∀f. f ⋓ f ≡ f. +corec lemma sor_idem: ∀f. f ⋓ f ≘ f. #f cases (pn_split f) * #g #H [ @(sor_pp … H H H) | @(sor_nn … H H H) ] -H // qed. -corec lemma sor_sym: ∀f1,f2,f. f1 ⋓ f2 ≡ f → f2 ⋓ f1 ≡ f. +corec lemma sor_comm: ∀f1,f2,f. f1 ⋓ f2 ≘ f → f2 ⋓ f1 ≘ f. #f1 #f2 #f * -f1 -f2 -f #f1 #f2 #f #g1 #g2 #g #Hf * * * -g1 -g2 -g [ @sor_pp | @sor_pn | @sor_np | @sor_nn ] /2 width=7 by/ qed-. -(* Properies on identity ****************************************************) +(* Properties with tail *****************************************************) + +lemma sor_tl: ∀f1,f2,f. f1 ⋓ f2 ≘ f → ⫱f1 ⋓ ⫱f2 ≘ ⫱f. +#f1 cases (pn_split f1) * #g1 #H1 +#f2 cases (pn_split f2) * #g2 #H2 +#f #Hf +[ cases (sor_inv_ppx … Hf … H1 H2) +| cases (sor_inv_pnx … Hf … H1 H2) +| cases (sor_inv_npx … Hf … H1 H2) +| cases (sor_inv_nnx … Hf … H1 H2) +] -Hf #g #Hg #H destruct // +qed. -corec lemma sor_isid_sn: ∀f1. 𝐈⦃f1⦄ → ∀f2. f1 ⋓ f2 ≡ f2. +lemma sor_xxn_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). +#g1 #g2 #g #H #f #H0 elim (sor_inv_xxn … H … H0) -H -H0 * +/3 width=5 by ex3_2_intro, or_introl, or_intror/ +qed-. + +lemma sor_xnx_tl: ∀g1,g2,g. g1 ⋓ g2 ≘ g → ∀f2. ↑f2 = g2 → + ∃∃f1,f. f1 ⋓ f2 ≘ f & ⫱g1 = f1 & ↑f = g. +#g1 elim (pn_split g1) * #f1 #H1 #g2 #g #H #f2 #H2 +[ elim (sor_inv_pnx … H … H1 H2) | elim (sor_inv_nnx … H … H1 H2) ] -g2 +/3 width=5 by ex3_2_intro/ +qed-. + +lemma sor_nxx_tl: ∀g1,g2,g. g1 ⋓ g2 ≘ g → ∀f1. ↑f1 = g1 → + ∃∃f2,f. f1 ⋓ f2 ≘ f & ⫱g2 = f2 & ↑f = g. +#g1 #g2 elim (pn_split g2) * #f2 #H2 #g #H #f1 #H1 +[ elim (sor_inv_npx … H … H1 H2) | elim (sor_inv_nnx … H … H1 H2) ] -g1 +/3 width=5 by ex3_2_intro/ +qed-. + +(* Properties with iterated tail ********************************************) + +lemma sor_tls: ∀f1,f2,f. f1 ⋓ f2 ≘ f → + ∀n. ⫱*[n]f1 ⋓ ⫱*[n]f2 ≘ ⫱*[n]f. +#f1 #f2 #f #Hf #n elim n -n /2 width=1 by sor_tl/ +qed. + +(* Properies with test for identity *****************************************) + +corec lemma sor_isid_sn: ∀f1. 𝐈❪f1❫ → ∀f2. f1 ⋓ f2 ≘ f2. #f1 * -f1 #f1 #g1 #Hf1 #H1 #f2 cases (pn_split f2) * /3 width=7 by sor_pp, sor_pn/ qed. -corec lemma sor_isid_dx: ∀f2. 𝐈⦃f2⦄ → ∀f1. f1 ⋓ f2 ≡ f1. +corec lemma sor_isid_dx: ∀f2. 𝐈❪f2❫ → ∀f1. f1 ⋓ f2 ≘ f1. #f2 * -f2 #f2 #g2 #Hf2 #H2 #f1 cases (pn_split f1) * /3 width=7 by sor_pp, sor_np/ qed. -(* Properties on finite colength assignment *********************************) +lemma sor_isid: ∀f1,f2,f. 𝐈❪f1❫ → 𝐈❪f2❫ → 𝐈❪f❫ → f1 ⋓ f2 ≘ f. +/4 width=3 by sor_eq_repl_back2, sor_eq_repl_back1, isid_inv_eq_repl/ qed. + +(* Inversion lemmas with tail ***********************************************) + +lemma sor_inv_tl_sn: ∀f1,f2,f. ⫱f1 ⋓ f2 ≘ f → f1 ⋓ ↑f2 ≘ ↑f. +#f1 #f2 #f elim (pn_split f1) * +#g1 #H destruct /2 width=7 by sor_pn, sor_nn/ +qed-. + +lemma sor_inv_tl_dx: ∀f1,f2,f. f1 ⋓ ⫱f2 ≘ f → ↑f1 ⋓ f2 ≘ ↑f. +#f1 #f2 #f elim (pn_split f2) * +#g2 #H destruct /2 width=7 by sor_np, sor_nn/ +qed-. + +(* Inversion lemmas with test for identity **********************************) -lemma sor_fcla_ex: ∀f1,n1. 𝐂⦃f1⦄ ≡ n1 → ∀f2,n2. 𝐂⦃f2⦄ ≡ n2 → - ∃∃f,n. f1 ⋓ f2 ≡ f & 𝐂⦃f⦄ ≡ n & (n1 ∨ n2) ≤ n & n ≤ n1 + n2. +lemma sor_isid_inv_sn: ∀f1,f2,f. f1 ⋓ f2 ≘ f → 𝐈❪f1❫ → f2 ≡ f. +/3 width=4 by sor_isid_sn, sor_mono/ +qed-. + +lemma sor_isid_inv_dx: ∀f1,f2,f. f1 ⋓ f2 ≘ f → 𝐈❪f2❫ → f1 ≡ f. +/3 width=4 by sor_isid_dx, sor_mono/ +qed-. + +corec lemma sor_fwd_isid1: ∀f1,f2,f. f1 ⋓ f2 ≘ f → 𝐈❪f❫ → 𝐈❪f1❫. +#f1 #f2 #f * -f1 -f2 -f +#f1 #f2 #f #g1 #g2 #g #Hf #H1 #H2 #H #Hg +[ /4 width=6 by isid_inv_push, isid_push/ ] +cases (isid_inv_next … Hg … H) +qed-. + +corec lemma sor_fwd_isid2: ∀f1,f2,f. f1 ⋓ f2 ≘ f → 𝐈❪f❫ → 𝐈❪f2❫. +#f1 #f2 #f * -f1 -f2 -f +#f1 #f2 #f #g1 #g2 #g #Hf #H1 #H2 #H #Hg +[ /4 width=6 by isid_inv_push, isid_push/ ] +cases (isid_inv_next … Hg … H) +qed-. + +lemma sor_inv_isid3: ∀f1,f2,f. f1 ⋓ f2 ≘ f → 𝐈❪f❫ → 𝐈❪f1❫ ∧ 𝐈❪f2❫. +/3 width=4 by sor_fwd_isid2, sor_fwd_isid1, conj/ qed-. + +(* Properties with finite colength assignment *******************************) + +lemma sor_fcla_ex: ∀f1,n1. 𝐂❪f1❫ ≘ n1 → ∀f2,n2. 𝐂❪f2❫ ≘ n2 → + ∃∃f,n. f1 ⋓ f2 ≘ f & 𝐂❪f❫ ≘ n & (n1 ∨ n2) ≤ n & n ≤ n1 + n2. #f1 #n1 #Hf1 elim Hf1 -f1 -n1 /3 width=6 by sor_isid_sn, ex4_2_intro/ #f1 #n1 #Hf1 #IH #f2 #n2 * -f2 -n2 /3 width=6 by fcla_push, fcla_next, ex4_2_intro, sor_isid_dx/ -#f2 #n2 #Hf2 elim (IH … Hf2) -IH -Hf2 -Hf1 -[ /3 width=7 by fcla_push, sor_pp, ex4_2_intro/ -| /3 width=7 by fcla_next, sor_pn, max_S2_le_S, le_S_S, ex4_2_intro/ -| /3 width=7 by fcla_next, sor_np, max_S1_le_S, le_S_S, ex4_2_intro/ +#f2 #n2 #Hf2 elim (IH … Hf2) -IH -Hf2 -Hf1 [2,4: #f #n