From 26d2ecb945a881c61d03f3c259996374209f5d7f Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Wed, 17 Jan 2018 20:31:56 +0100 Subject: [PATCH] update in basic_2 yet another definition of lveq allows to prove a missing lemma in lfxs_fle --- .gitignore | 1 + .../lambdadelta/basic_2/static/fle.ma | 8 +- .../lambdadelta/basic_2/static/fle_drops.ma | 1 - .../lambdadelta/basic_2/static/fle_fle.ma | 9 +- .../lambdadelta/basic_2/static/fle_fqup.ma | 3 +- .../lambdadelta/basic_2/static/lfxs_fle.ma | 23 +-- .../lambdadelta/basic_2/syntax/lenv_length.ma | 2 - .../lambdadelta/basic_2/syntax/lveq.ma | 177 ++++++++++-------- .../lambdadelta/basic_2/syntax/lveq_length.ma | 54 +++--- .../lambdadelta/basic_2/syntax/lveq_lveq.ma | 141 +++----------- .../lambdadelta/basic_2/web/basic_2_src.tbl | 11 +- 11 files changed, 175 insertions(+), 255 deletions(-) diff --git a/.gitignore b/.gitignore index 7ab34bb92..45d3126b7 100644 --- a/.gitignore +++ b/.gitignore @@ -66,6 +66,7 @@ helm/www/lambdadelta/xslt/core.xsl helm/www/lambdadelta/xslt/chc_45.xsl helm/www/lambdadelta/xslt/xhtbl.xsl +matita/matita/contribs/lambdadelta/nodes matita/matita/contribs/lambdadelta/token matita/matita/contribs/lambdadelta/2A1 matita/matita/contribs/lambdadelta/apps_2/notation diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/fle.ma b/matita/matita/contribs/lambdadelta/basic_2/static/fle.ma index fdc03a5e2..95d12b132 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/fle.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/fle.ma @@ -29,14 +29,10 @@ interpretation "free variables inclusion (restricted closure)" (* Basic properties *********************************************************) lemma fle_sort: ∀L,s1,s2. ⦃L, ⋆s1⦄ ⊆ ⦃L, ⋆s2⦄. -#L elim (lveq_refl L) -/3 width=8 by frees_sort, sle_refl, ex4_4_intro/ -qed. +/3 width=8 by frees_sort, sle_refl, ex4_4_intro/ qed. lemma fle_gref: ∀L,l1,l2. ⦃L, §l1⦄ ⊆ ⦃L, §l2⦄. -#L elim (lveq_refl L) -/3 width=8 by frees_gref, sle_refl, ex4_4_intro/ -qed. +/3 width=8 by frees_gref, sle_refl, ex4_4_intro/ qed. (* Basic inversion lemmas ***************************************************) (* diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/fle_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/static/fle_drops.ma index ff0b8872d..4ce02b60b 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/fle_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/fle_drops.ma @@ -26,6 +26,5 @@ elim (frees_total L V) #f1 #Hf1 elim (frees_total L T) #f2 #Hf2 elim (sor_isfin_ex f1 f2) /3 width=3 by frees_fwd_isfin, isfin_tl/ #f #Hf #_ lapply (sor_inv_sle_dx … Hf) #Hf0 -elim (lveq_refl L) #n #HL /6 width=8 by frees_lifts_SO, frees_bind, drops_refl, drops_drop, sle_tls, ex4_4_intro/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/fle_fle.ma b/matita/matita/contribs/lambdadelta/basic_2/static/fle_fle.ma index 42c403412..4301456f5 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/fle_fle.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/fle_fle.ma @@ -38,11 +38,12 @@ theorem fle_trans: bi_transitive … fle. /5 width=8 by frees_mono, sle_trans, sle_eq_repl_back2, ex3_2_intro/ qed-. *) -theorem fle_bind_sn: ∀L1,L2,V1,T1,T. ⦃L1, V1⦄ ⊆ ⦃L2, T⦄ → ⦃L1.ⓧ, T1⦄ ⊆ ⦃L2, T⦄ → - ∀p,I. ⦃L1, ⓑ{p,I}V1.T1⦄ ⊆ ⦃L2, T⦄. -#L1 #L2 #V1 #T1 #T * #n1 #x #f1 #g #Hf1 #Hg #H1n1 #H2n1 #H #p #I +theorem fle_bind_sn_ge: ∀L1,L2. |L2| ≤ |L1| → + ∀V1,T1,T. ⦃L1, V1⦄ ⊆ ⦃L2, T⦄ → ⦃L1.ⓧ, T1⦄ ⊆ ⦃L2, T⦄ → + ∀p,I. ⦃L1, ⓑ{p,I}V1.T1⦄ ⊆ ⦃L2, T⦄. +#L1 #L2 #HL #V1 #T1 #T * #n1 #x #f1 #g #Hf1 #Hg #H1n1 #H2n1 #H #p #I elim (fle_frees_trans … H … Hg) -H #n2 #n #f2 #Hf2 #H1n2 #H2n2 -elim (lveq_inj_void_sn … H1n1 … H1n2) -H1n2 #H1 #H2 destruct +elim (lveq_inj_void_sn_ge … H1n1 … H1n2) -H1n2 // #H1 #H2 #H3 destruct elim (sor_isfin_ex f1 (⫱f2)) /3 width=3 by frees_fwd_isfin, isfin_tl/ #f #Hf #_ (length_inv_zero_sn … H) -Y2 /2 width=3 by lveq_atom, ex_intro/ -| #K1 * [ * | #I1 #V1 ] #IH #Y2 #H - elim (length_inv_succ_sn … H) -H * [1,3: * |*: #I2 #V2 ] #K2 #HK #H destruct - elim (IH … HK) -IH -HK #n #HK - /4 width=3 by lveq_pair_sn, lveq_pair_dx, lveq_void_sn, lveq_void_dx, ex_intro/ +| #K1 #I1 #IH #Y2 #H + elim (length_inv_succ_sn … H) -H #I2 #K2 #HK #H destruct + /3 width=1 by lveq_bind/ ] -qed-. +qed. (* Forward lemmas with length for local environments ************************) @@ -40,31 +41,36 @@ lemma lveq_fwd_length_le_dx: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 → n2 ≤ |L qed-. lemma lveq_fwd_length: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 → - |L1| + n2 = |L2| + n1. + ∧∧ |L1|-|L2| = n1 & |L2|-|L1| = n2. +#L1 #L2 #n1 #n2 #H elim H -L1 -L2 -n1 -n2 /2 width=1 by conj/ +#K1 #K2 #n #_ * #H1 #H2 >length_bind /3 width=1 by minus_Sn_m, conj/ +qed-. + +lemma lveq_inj_length: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 → + |L1| = |L2| → ∧∧ 0 = n1 & 0 = n2. +#L1 #L2 #n1 #n2 #H #HL +elim (lveq_fwd_length … H) -H +>HL -HL /2 width=1 by conj/ +qed-. + +lemma lveq_fwd_length_plus: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 → + |L1| + n2 = |L2| + n1. #L1 #L2 #n1 #n2 #H elim H -L1 -L2 -n1 -n2 normalize /2 width=2 by injective_plus_r/ qed-. -lemma lveq_fwd_length_eq: ∀L1,L2,n. L1 ≋ⓧ*[n, n] L2 → |L1| = |L2|. -/3 width=2 by lveq_fwd_length, injective_plus_l/ qed-. +lemma lveq_fwd_length_eq: ∀L1,L2. L1 ≋ⓧ*[0, 0] L2 → |L1| = |L2|. +/3 width=2 by lveq_fwd_length_plus, injective_plus_l/ qed-. lemma lveq_fwd_length_minus: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 → |L1| - n1 = |L2| - n2. -/3 width=3 by lveq_fwd_length, lveq_fwd_length_le_dx, lveq_fwd_length_le_sn, plus_to_minus_2/ qed-. - -lemma lveq_inj_length: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 → - |L1| = |L2| → n1 = n2. -#L1 #L2 #n1 #n2 #H #HL12 -lapply (lveq_fwd_length … H) -H #H -/2 width=2 by injective_plus_l/ -qed-. +/3 width=3 by lveq_fwd_length_plus, lveq_fwd_length_le_dx, lveq_fwd_length_le_sn, plus_to_minus_2/ qed-. lemma lveq_fwd_abst_bind_length_le: ∀I1,I2,L1,L2,V1,n1,n2. L1.ⓑ{I1}V1 ≋ⓧ*[n1, n2] L2.ⓘ{I2} → |L1| ≤ |L2|. #I1 #I2 #L1 #L2 #V1 #n1 #n2 #HL lapply (lveq_fwd_pair_sn … HL) #H destruct -lapply (lveq_fwd_length … HL) -HL >length_bind >length_bind #H -/2 width=1 by monotonic_pred/ +elim (lveq_fwd_length … HL) -HL >length_bind >length_bind // qed-. lemma lveq_fwd_bind_abst_length_le: ∀I1,I2,L1,L2,V2,n1,n2. @@ -74,16 +80,16 @@ lemma lveq_fwd_bind_abst_length_le: ∀I1,I2,L1,L2,V2,n1,n2. (* Inversion lemmas with length for local environments **********************) lemma lveq_inv_void_dx_length: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2.ⓧ → |L1| ≤ |L2| → - ∃∃m2. L1 ≋ ⓧ*[n1, m2] L2 & n2 = ⫯m2 & n1 ≤ m2. + ∃∃m2. L1 ≋ ⓧ*[n1, m2] L2 & 0 = n1 & ⫯m2 = n2. #L1 #L2 #n1 #n2 #H #HL12 -lapply (lveq_fwd_length … H) normalize >plus_n_Sm #H0 +lapply (lveq_fwd_length_plus … H) normalize >plus_n_Sm #H0 lapply (plus2_inv_le_sn … H0 HL12) -H0 -HL12 #H0 -elim (le_inv_S1 … H0) -H0 #m2 #Hm2 #H0 destruct -/3 width=4 by lveq_inv_void_succ_dx, ex3_intro/ +elim (le_inv_S1 … H0) -H0 #m2 #_ #H0 destruct +elim (lveq_inv_void_succ_dx … H) -H /2 width=3 by ex3_intro/ qed-. lemma lveq_inv_void_sn_length: ∀L1,L2,n1,n2. L1.ⓧ ≋ⓧ*[n1, n2] L2 → |L2| ≤ |L1| → - ∃∃m1. L1 ≋ ⓧ*[m1, n2] L2 & n1 = ⫯m1 & n2 ≤ m1. + ∃∃m1. L1 ≋ ⓧ*[m1, n2] L2 & ⫯m1 = n1 & 0 = n2. #L1 #L2 #n1 #n2 #H #HL lapply (lveq_sym … H) -H #H elim (lveq_inv_void_dx_length … H HL) -H -HL diff --git a/matita/matita/contribs/lambdadelta/basic_2/syntax/lveq_lveq.ma b/matita/matita/contribs/lambdadelta/basic_2/syntax/lveq_lveq.ma index 3f64c33b0..3d3417dba 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/syntax/lveq_lveq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/syntax/lveq_lveq.ma @@ -16,131 +16,38 @@ include "basic_2/syntax/lveq_length.ma". (* EQUIVALENCE FOR LOCAL ENVIRONMENTS UP TO EXCLUSION BINDERS ***************) -(* Main forward lemmas ******************************************************) - -theorem lveq_fwd_inj_succ_zero: ∀L1,L2,n1. L1 ≋ⓧ*[⫯n1, 0] L2 → - ∀m1,m2. L1 ≋ⓧ*[m1, m2] L2 → ∃x1. ⫯x1 = m1. -#L1 #L2 #n1 #Hn #m1 #m2 #Hm -lapply (lveq_fwd_length … Hn) -Hn Hn >associative_plus -Hn #Hm -lapply (injective_plus_r … Hm) -Hm -length_bind >eq_minus_S_pred >(eq_minus_O … HL) +/3 width=4 by plus_minus, and3_intro/ +qed-. -theorem lveq_inj_void_dx: ∀K1,K2,n1,n2. K1 ≋ⓧ*[n1, n2] K2 → - ∀m1,m2. K1 ≋ⓧ*[m1, m2] K2.ⓧ → - ∧∧ n1 = m1 & ⫯n2 = m2. -/3 width=4 by lveq_inj, lveq_void_dx/ qed-. +theorem lveq_inj_void_dx_le: ∀K1,K2. |K1| ≤ |K2| → + ∀n1,n2. K1 ≋ⓧ*[n1, n2] K2 → + ∀m1,m2. K1 ≋ⓧ*[m1, m2] K2.ⓧ → + ∧∧ ⫯n2 = m2 & 0 = m1 & 0 = n1. +/3 width=5 by lveq_inj_void_sn_ge, lveq_sym/ qed-. (* auto: 2x lveq_sym *) diff --git a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl index 0dd04cb2c..618d4550e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl +++ b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl @@ -239,16 +239,15 @@ table { class "red" [ { "syntax" * } { [ { "equivalence up to exclusion binders" * } { - [ [ "" ] "voids ( ⓧ*[?]? ≋ ⓧ*[?]? )" "voids_length" + "voids_voids" * ] + [ [ "for lenvs" ] "lveq ( ? ≋ⓧ*[?,?] ? )" "lveq_length" + "lveq_lveq" * ] } ] - - [ { "append for local environments" * } { - [ [ "" ] "append ( ? @@ ? )" "append_length" * ] + [ { "append" * } { + [ [ "for lenvs" ] "append ( ? @@ ? )" "append_length" * ] } ] - [ { "head equivalence for terms" * } { - [ [ "" ] "theq ( ? ⩳[?,?] ? )" "theq_simple" + "theq_tdeq" + "theq_theq" + "theq_simple_vector" * ] + [ { "head equivalence" * } { + [ [ "for terms" ] "theq ( ? ⩳[?,?] ? )" "theq_simple" + "theq_tdeq" + "theq_theq" + "theq_simple_vector" * ] } ] [ { "degree-based equivalence" * } { -- 2.39.2