X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frelocation%2Flsubr.ma;h=e96683a80e96fc536957699695a2065bc013b7e4;hb=8676e97d61b676fac6b740f6e10503672e992c00;hp=faf254e36e6aabeb10892a6e0b4b2a90ff4b7bca;hpb=3e9d72c26091f0e157a024ea9bd6f95a95729860;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lsubr.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lsubr.ma index faf254e36..e96683a80 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lsubr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lsubr.ma @@ -30,7 +30,7 @@ interpretation (* Basic properties *********************************************************) lemma lsubr_refl: ∀L. L ⊑ L. -#L elim L -L // /2 width=1/ +#L elim L -L /2 width=1 by lsubr_sort, lsubr_bind/ qed. (* Basic inversion lemmas ***************************************************) @@ -48,8 +48,8 @@ lemma lsubr_inv_atom1: ∀L2. ⋆ ⊑ L2 → L2 = ⋆. fact lsubr_inv_abst1_aux: ∀L1,L2. L1 ⊑ L2 → ∀K1,W. L1 = K1.ⓛW → L2 = ⋆ ∨ ∃∃K2. K1 ⊑ K2 & L2 = K2.ⓛW. #L1 #L2 * -L1 -L2 -[ #L #K1 #W #H destruct /2 width=1/ -| #I #L1 #L2 #V #HL12 #K1 #W #H destruct /3 width=3/ +[ #L #K1 #W #H destruct /2 width=1 by or_introl/ +| #I #L1 #L2 #V #HL12 #K1 #W #H destruct /3 width=3 by ex2_intro, or_intror/ | #L1 #L2 #V1 #V2 #_ #K1 #W #H destruct ] qed-. @@ -62,7 +62,7 @@ fact lsubr_inv_abbr2_aux: ∀L1,L2. L1 ⊑ L2 → ∀K2,W. L2 = K2.ⓓW → ∃∃K1. K1 ⊑ K2 & L1 = K1.ⓓW. #L1 #L2 * -L1 -L2 [ #L #K2 #W #H destruct -| #I #L1 #L2 #V #HL12 #K2 #W #H destruct /2 width=3/ +| #I #L1 #L2 #V #HL12 #K2 #W #H destruct /2 width=3 by ex2_intro/ | #L1 #L2 #V1 #V2 #_ #K2 #W #H destruct ] qed-. @@ -74,32 +74,34 @@ lemma lsubr_inv_abbr2: ∀L1,K2,W. L1 ⊑ K2.ⓓW → (* Basic forward lemmas *****************************************************) lemma lsubr_fwd_length: ∀L1,L2. L1 ⊑ L2 → |L2| ≤ |L1|. -#L1 #L2 #H elim H -L1 -L2 // /2 width=1/ +#L1 #L2 #H elim H -L1 -L2 /2 width=1 by monotonic_le_plus_l/ qed-. lemma lsubr_fwd_ldrop2_bind: ∀L1,L2. L1 ⊑ L2 → - ∀I,K2,W,i. ⇩[0, i] L2 ≡ K2.ⓑ{I}W → - (∃∃K1. K1 ⊑ K2 & ⇩[0, i] L1 ≡ K1.ⓑ{I}W) ∨ - ∃∃K1,V. K1 ⊑ K2 & ⇩[0, i] L1 ≡ K1.ⓓⓝW.V & I = Abst. + ∀I,K2,W,s,i. ⇩[s, 0, i] L2 ≡ K2.ⓑ{I}W → + (∃∃K1. K1 ⊑ K2 & ⇩[s, 0, i] L1 ≡ K1.ⓑ{I}W) ∨ + ∃∃K1,V. K1 ⊑ K2 & ⇩[s, 0, i] L1 ≡ K1.ⓓⓝW.V & I = Abst. #L1 #L2 #H elim H -L1 -L2 -[ #L #I #K2 #W #i #H +[ #L #I #K2 #W #s #i #H elim (ldrop_inv_atom1 … H) -H #H destruct -| #J #L1 #L2 #V #HL12 #IHL12 #I #K2 #W #i #H +| #J #L1 #L2 #V #HL12 #IHL12 #I #K2 #W #s #i #H elim (ldrop_inv_O1_pair1 … H) -H * #Hi #HLK2 destruct [ -IHL12 | -HL12 ] - [ /3 width=3/ - | elim (IHL12 … HLK2) -IHL12 -HLK2 * /4 width=3/ /4 width=4/ + [ /3 width=3 by ldrop_pair, ex2_intro, or_introl/ + | elim (IHL12 … HLK2) -IHL12 -HLK2 * + /4 width=4 by ldrop_drop_lt, ex3_2_intro, ex2_intro, or_introl, or_intror/ ] -| #L1 #L2 #V1 #V2 #HL12 #IHL12 #I #K2 #W #i #H +| #L1 #L2 #V1 #V2 #HL12 #IHL12 #I #K2 #W #s #i #H elim (ldrop_inv_O1_pair1 … H) -H * #Hi #HLK2 destruct [ -IHL12 | -HL12 ] - [ /3 width=4/ - | elim (IHL12 … HLK2) -IHL12 -HLK2 * /4 width=3/ /4 width=4/ + [ /3 width=4 by ldrop_pair, ex3_2_intro, or_intror/ + | elim (IHL12 … HLK2) -IHL12 -HLK2 * + /4 width=4 by ldrop_drop_lt, ex3_2_intro, ex2_intro, or_introl, or_intror/ ] ] qed-. lemma lsubr_fwd_ldrop2_abbr: ∀L1,L2. L1 ⊑ L2 → - ∀K2,V,i. ⇩[0, i] L2 ≡ K2.ⓓV → - ∃∃K1. K1 ⊑ K2 & ⇩[0, i] L1 ≡ K1.ⓓV. -#L1 #L2 #HL12 #K2 #V #i #HLK2 elim (lsubr_fwd_ldrop2_bind … HL12 … HLK2) -L2 // * + ∀K2,V,s,i. ⇩[s, 0, i] L2 ≡ K2.ⓓV → + ∃∃K1. K1 ⊑ K2 & ⇩[s, 0, i] L1 ≡ K1.ⓓV. +#L1 #L2 #HL12 #K2 #V #s #i #HLK2 elim (lsubr_fwd_ldrop2_bind … HL12 … HLK2) -L2 // * #K1 #W #_ #_ #H destruct qed-.