From: Ferruccio Guidi Date: Mon, 25 Jun 2018 22:54:48 +0000 (+0200) Subject: renaming in basic_2 X-Git-Tag: make_still_working~302 X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=commitdiff_plain;h=222044da28742b24584549ba86b1805a87def070 renaming in basic_2 + a notation problem solved --- diff --git a/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs.ma b/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs.ma new file mode 100644 index 000000000..e8bd0a2bd --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs.ma @@ -0,0 +1,195 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "ground_2/lib/star.ma". +include "basic_2/notation/relations/relationstar_4.ma". +include "basic_2/static/rex.ma". + +(* ITERATED EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ***) + +definition rexs (R): term → relation lenv ≝ CTC … (rex R). + +interpretation "iterated extension on referred entries (local environment)" + 'RelationStar R T L1 L2 = (rexs R T L1 L2). + +(* Basic properties *********************************************************) + +lemma rexs_step_dx: ∀R,L1,L,T. L1 ⪤*[R, T] L → + ∀L2. L ⪤[R, T] L2 → L1 ⪤*[R, T] L2. +#R #L1 #L2 #T #HL1 #L2 @step @HL1 (**) (* auto fails *) +qed-. + +lemma rexs_step_sn: ∀R,L1,L,T. L1 ⪤[R, T] L → + ∀L2. L ⪤*[R, T] L2 → L1 ⪤*[R, T] L2. +#R #L1 #L2 #T #HL1 #L2 @TC_strap @HL1 (**) (* auto fails *) +qed-. + +lemma rexs_atom: ∀R,I. ⋆ ⪤*[R, ⓪{I}] ⋆. +/2 width=1 by inj/ qed. + +lemma rexs_sort: ∀R,I,L1,L2,V1,V2,s. + L1 ⪤*[R, ⋆s] L2 → L1.ⓑ{I}V1 ⪤*[R, ⋆s] L2.ⓑ{I}V2. +#R #I #L1 #L2 #V1 #V2 #s #H elim H -L2 +/3 width=4 by rex_sort, rexs_step_dx, inj/ +qed. + +lemma rexs_pair: ∀R. (∀L. reflexive … (R L)) → + ∀I,L1,L2,V. L1 ⪤*[R, V] L2 → + L1.ⓑ{I}V ⪤*[R, #0] L2.ⓑ{I}V. +#R #HR #I #L1 #L2 #V #H elim H -L2 +/3 width=5 by rex_pair, rexs_step_dx, inj/ +qed. + +lemma rexs_unit: ∀R,f,I,L1,L2. 𝐈⦃f⦄ → L1 ⪤[cext2 R, cfull, f] L2 → + L1.ⓤ{I} ⪤*[R, #0] L2.ⓤ{I}. +/3 width=3 by rex_unit, inj/ qed. + +lemma rexs_lref: ∀R,I,L1,L2,V1,V2,i. + L1 ⪤*[R, #i] L2 → L1.ⓑ{I}V1 ⪤*[R, #↑i] L2.ⓑ{I}V2. +#R #I #L1 #L2 #V1 #V2 #i #H elim H -L2 +/3 width=4 by rex_lref, rexs_step_dx, inj/ +qed. + +lemma rexs_gref: ∀R,I,L1,L2,V1,V2,l. + L1 ⪤*[R, §l] L2 → L1.ⓑ{I}V1 ⪤*[R, §l] L2.ⓑ{I}V2. +#R #I #L1 #L2 #V1 #V2 #l #H elim H -L2 +/3 width=4 by rex_gref, rexs_step_dx, inj/ +qed. + +lemma rexs_co: ∀R1,R2. (∀L,T1,T2. R1 L T1 T2 → R2 L T1 T2) → + ∀L1,L2,T. L1 ⪤*[R1, T] L2 → L1 ⪤*[R2, T] L2. +#R1 #R2 #HR #L1 #L2 #T #H elim H -L2 +/4 width=5 by rex_co, rexs_step_dx, inj/ +qed-. + +(* Basic inversion lemmas ***************************************************) + +(* Basic_2A1: uses: TC_lpx_sn_inv_atom1 *) +lemma rexs_inv_atom_sn: ∀R,I,Y2. ⋆ ⪤*[R, ⓪{I}] Y2 → Y2 = ⋆. +#R #I #Y2 #H elim H -Y2 /3 width=3 by inj, rex_inv_atom_sn/ +qed-. + +(* Basic_2A1: uses: TC_lpx_sn_inv_atom2 *) +lemma rexs_inv_atom_dx: ∀R,I,Y1. Y1 ⪤*[R, ⓪{I}] ⋆ → Y1 = ⋆. +#R #I #Y1 #H @(TC_ind_dx ??????? H) -Y1 +/3 width=3 by inj, rex_inv_atom_dx/ +qed-. + +lemma rexs_inv_sort: ∀R,Y1,Y2,s. Y1 ⪤*[R, ⋆s] Y2 → + ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆ + | ∃∃I1,I2,L1,L2. L1 ⪤*[R, ⋆s] L2 & + Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. +#R #Y1 #Y2 #s #H elim H -Y2 +[ #Y2 #H elim (rex_inv_sort … H) -H * + /4 width=8 by ex3_4_intro, inj, or_introl, or_intror, conj/ +| #Y #Y2 #_ #H elim (rex_inv_sort … H) -H * + [ #H #H2 * * /3 width=7 by ex3_4_intro, or_introl, or_intror, conj/ + | #I #I2 #L #L2 #HL2 #H #H2 * * + [ #H1 #H0 destruct + | #I1 #I0 #L1 #L0 #HL10 #H1 #H0 destruct + /4 width=7 by ex3_4_intro, rexs_step_dx, or_intror/ + ] + ] +] +qed-. + +lemma rexs_inv_gref: ∀R,Y1,Y2,l. Y1 ⪤*[R, §l] Y2 → + ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆ + | ∃∃I1,I2,L1,L2. L1 ⪤*[R, §l] L2 & + Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. +#R #Y1 #Y2 #l #H elim H -Y2 +[ #Y2 #H elim (rex_inv_gref … H) -H * + /4 width=8 by ex3_4_intro, inj, or_introl, or_intror, conj/ +| #Y #Y2 #_ #H elim (rex_inv_gref … H) -H * + [ #H #H2 * * /3 width=7 by ex3_4_intro, or_introl, or_intror, conj/ + | #I #I2 #L #L2 #HL2 #H #H2 * * + [ #H1 #H0 destruct + | #I1 #I0 #L1 #L0 #HL10 #H1 #H0 destruct + /4 width=7 by ex3_4_intro, rexs_step_dx, or_intror/ + ] + ] +] +qed-. + +lemma rexs_inv_bind: ∀R. (∀L. reflexive … (R L)) → + ∀p,I,L1,L2,V,T. L1 ⪤*[R, ⓑ{p,I}V.T] L2 → + ∧∧ L1 ⪤*[R, V] L2 & L1.ⓑ{I}V ⪤*[R, T] L2.ⓑ{I}V. +#R #HR #p #I #L1 #L2 #V #T #H elim H -L2 +[ #L2 #H elim (rex_inv_bind … V ? H) -H /3 width=1 by inj, conj/ +| #L #L2 #_ #H * elim (rex_inv_bind … V ? H) -H /3 width=3 by rexs_step_dx, conj/ +] +qed-. + +lemma rexs_inv_flat: ∀R,I,L1,L2,V,T. L1 ⪤*[R, ⓕ{I}V.T] L2 → + ∧∧ L1 ⪤*[R, V] L2 & L1 ⪤*[R, T] L2. +#R #I #L1 #L2 #V #T #H elim H -L2 +[ #L2 #H elim (rex_inv_flat … H) -H /3 width=1 by inj, conj/ +| #L #L2 #_ #H * elim (rex_inv_flat … H) -H /3 width=3 by rexs_step_dx, conj/ +] +qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma rexs_inv_sort_bind_sn: ∀R,I1,Y2,L1,s. L1.ⓘ{I1} ⪤*[R, ⋆s] Y2 → + ∃∃I2,L2. L1 ⪤*[R, ⋆s] L2 & Y2 = L2.ⓘ{I2}. +#R #I1 #Y2 #L1 #s #H elim (rexs_inv_sort … H) -H * +[ #H destruct +| #Z #I2 #Y1 #L2 #Hs #H1 #H2 destruct /2 width=4 by ex2_2_intro/ +] +qed-. + +lemma rexs_inv_sort_bind_dx: ∀R,I2,Y1,L2,s. Y1 ⪤*[R, ⋆s] L2.ⓘ{I2} → + ∃∃I1,L1. L1 ⪤*[R, ⋆s] L2 & Y1 = L1.ⓘ{I1}. +#R #I2 #Y1 #L2 #s #H elim (rexs_inv_sort … H) -H * +[ #_ #H destruct +| #I1 #Z #L1 #Y2 #Hs #H1 #H2 destruct /2 width=4 by ex2_2_intro/ +] +qed-. + +lemma rexs_inv_gref_bind_sn: ∀R,I1,Y2,L1,l. L1.ⓘ{I1} ⪤*[R, §l] Y2 → + ∃∃I2,L2. L1 ⪤*[R, §l] L2 & Y2 = L2.ⓘ{I2}. +#R #I1 #Y2 #L1 #l #H elim (rexs_inv_gref … H) -H * +[ #H destruct +| #Z #I2 #Y1 #L2 #Hl #H1 #H2 destruct /2 width=4 by ex2_2_intro/ +] +qed-. + +lemma rexs_inv_gref_bind_dx: ∀R,I2,Y1,L2,l. Y1 ⪤*[R, §l] L2.ⓘ{I2} → + ∃∃I1,L1. L1 ⪤*[R, §l] L2 & Y1 = L1.ⓘ{I1}. +#R #I2 #Y1 #L2 #l #H elim (rexs_inv_gref … H) -H * +[ #_ #H destruct +| #I1 #Z #L1 #Y2 #Hl #H1 #H2 destruct /2 width=4 by ex2_2_intro/ +] +qed-. + +(* Basic forward lemmas *****************************************************) + +lemma rexs_fwd_pair_sn: ∀R,I,L1,L2,V,T. L1 ⪤*[R, ②{I}V.T] L2 → L1 ⪤*[R, V] L2. +#R #I #L1 #L2 #V #T #H elim H -L2 +/3 width=5 by rex_fwd_pair_sn, rexs_step_dx, inj/ +qed-. + +lemma rexs_fwd_bind_dx: ∀R. (∀L. reflexive … (R L)) → + ∀p,I,L1,L2,V,T. L1 ⪤*[R, ⓑ{p,I}V.T] L2 → + L1.ⓑ{I}V ⪤*[R, T] L2.ⓑ{I}V. +#R #HR #p #I #L1 #L2 #V #T #H elim (rexs_inv_bind … H) -H // +qed-. + +lemma rexs_fwd_flat_dx: ∀R,I,L1,L2,V,T. L1 ⪤*[R, ⓕ{I}V.T] L2 → L1 ⪤*[R, T] L2. +#R #I #L1 #L2 #V #T #H elim (rexs_inv_flat … H) -H // +qed-. + +(* Basic_2A1: removed theorems 2: + TC_lpx_sn_inv_pair1 TC_lpx_sn_inv_pair2 +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs_drops.ma new file mode 100644 index 000000000..3ace3be2e --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs_drops.ma @@ -0,0 +1,72 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/seq_seq.ma". +include "basic_2/static/rex_drops.ma". +include "basic_2/i_static/rexs.ma". + +(* ITERATED EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ***) + +definition tc_f_dedropable_sn: predicate (relation3 lenv term term) ≝ + λR. ∀b,f,L1,K1. ⬇*[b, f] L1 ≘ K1 → + ∀K2,T. K1 ⪤*[R, T] K2 → ∀U. ⬆*[f] T ≘ U → + ∃∃L2. L1 ⪤*[R, U] L2 & ⬇*[b, f] L2 ≘ K2 & L1 ≡[f] L2. + +definition tc_f_dropable_sn: predicate (relation3 lenv term term) ≝ + λR. ∀b,f,L1,K1. ⬇*[b, f] L1 ≘ K1 → 𝐔⦃f⦄ → + ∀L2,U. L1 ⪤*[R, U] L2 → ∀T. ⬆*[f] T ≘ U → + ∃∃K2. K1 ⪤*[R, T] K2 & ⬇*[b, f] L2 ≘ K2. + +definition tc_f_dropable_dx: predicate (relation3 lenv term term) ≝ + λR. ∀L1,L2,U. L1 ⪤*[R, U] L2 → + ∀b,f,K2. ⬇*[b, f] L2 ≘ K2 → 𝐔⦃f⦄ → ∀T. ⬆*[f] T ≘ U → + ∃∃K1. ⬇*[b, f] L1 ≘ K1 & K1 ⪤*[R, T] K2. + +(* Properties with generic slicing for local environments *******************) + +lemma dedropable_sn_CTC: ∀R. f_dedropable_sn R → tc_f_dedropable_sn R. +#R #HR #b #f #L1 #K1 #HLK1 #K2 #T #H elim H -K2 +[ #K2 #HK12 #U #HTU elim (HR … HLK1 … HK12 … HTU) -K1 -T -HR + /3 width=4 by ex3_intro, inj/ +| #K #K2 #_ #HK2 #IH #U #HTU -HLK1 + elim (IH … HTU) -IH #L #HL1 #HLK + elim (HR … HLK … HK2 … HTU) -K -T -HR + /3 width=6 by seq_trans, rexs_step_dx, ex3_intro/ +] +qed-. + +(* Inversion lemmas with generic slicing for local environments *************) + +lemma dropable_sn_CTC: ∀R. f_dropable_sn R → tc_f_dropable_sn R. +#R #HR #b #f #L1 #K1 #HLK1 #Hf #L2 #U #H elim H -L2 +[ #L2 #HL12 #T #HTU elim (HR … HLK1 … HL12 … HTU) -L1 -U -HR + /3 width=3 by inj, ex2_intro/ +| #L #L2 #_ #HL2 #IH #T #HTU -HLK1 + elim (IH … HTU) -IH #K #HK1 #HLK + elim (HR … HLK … HL2 … HTU) -L -U -HR + /3 width=3 by rexs_step_dx, ex2_intro/ +] +qed-. + +lemma dropable_dx_CTC: ∀R. f_dropable_dx R → tc_f_dropable_dx R. +#R #HR #L1 #L2 #U #H elim H -L2 +[ #L2 #HL12 #b #f #K2 #HLK2 #Hf #T #HTU + elim (HR … HL12 … HLK2 … HTU) -L2 -U -HR + /3 width=3 by inj, ex2_intro/ +| #L #L2 #_ #HL2 #IH #b #f #K2 #HLK2 #Hf #T #HTU + elim (HR … HL2 … HLK2 … HTU) -L2 -HR // #K #HLK #HK2 + elim (IH … HLK … HTU) -IH -L -U + /3 width=5 by rexs_step_dx, ex2_intro/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs_fqup.ma b/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs_fqup.ma new file mode 100644 index 000000000..ad5adc659 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs_fqup.ma @@ -0,0 +1,75 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rex_fqup.ma". +include "basic_2/i_static/rexs.ma". + +(* ITERATED EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ***) + +(* Advanced properties ******************************************************) + +lemma rexs_refl: ∀R. c_reflexive … R → + ∀T. reflexive … (rexs R T). +/3 width=1 by rex_refl, inj/ qed. + +(* Basic_2A1: uses: TC_lpx_sn_pair TC_lpx_sn_pair_refl *) +lemma rexs_pair_refl: ∀R. c_reflexive … R → + ∀L,V1,V2. CTC … R L V1 V2 → ∀I,T. L.ⓑ{I}V1 ⪤*[R, T] L.ⓑ{I}V2. +#R #HR #L #V1 #V2 #H elim H -V2 +/3 width=3 by rexs_step_dx, rex_pair_refl, inj/ +qed. + +lemma rexs_tc: ∀R,L1,L2,T,f. 𝐈⦃f⦄ → TC … (sex cfull (cext2 R) f) L1 L2 → + L1 ⪤*[R, T] L2. +#R #L1 #L2 #T #f #Hf #H elim H -L2 +[ elim (frees_total L1 T) | #L elim (frees_total L T) ] +/5 width=7 by sex_sdj, rexs_step_dx, sdj_isid_sn, inj, ex2_intro/ +qed. + +(* Advanced eliminators *****************************************************) + +lemma rexs_ind_sn: ∀R. c_reflexive … R → + ∀L1,T. ∀Q:predicate …. Q L1 → + (∀L,L2. L1 ⪤*[R, T] L → L ⪤[R, T] L2 → Q L → Q L2) → + ∀L2. L1 ⪤*[R, T] L2 → Q L2. +#R #HR #L1 #T #Q #HL1 #IHL1 #L2 #HL12 +@(TC_star_ind … HL1 IHL1 … HL12) /2 width=1 by rex_refl/ +qed-. + +lemma rexs_ind_dx: ∀R. c_reflexive … R → + ∀L2,T. ∀Q:predicate …. Q L2 → + (∀L1,L. L1 ⪤[R, T] L → L ⪤*[R, T] L2 → Q L → Q L1) → + ∀L1. L1 ⪤*[R, T] L2 → Q L1. +#R #HR #L2 #Q #HL2 #IHL2 #L1 #HL12 +@(TC_star_ind_dx … HL2 IHL2 … HL12) /2 width=4 by rex_refl/ +qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma rexs_inv_bind_void: ∀R. c_reflexive … R → + ∀p,I,L1,L2,V,T. L1 ⪤*[R, ⓑ{p,I}V.T] L2 → + ∧∧ L1 ⪤*[R, V] L2 & L1.ⓧ ⪤*[R, T] L2.ⓧ. +#R #HR #p #I #L1 #L2 #V #T #H @(rexs_ind_sn … HR … H) -L2 +[ /3 width=1 by rexs_refl, conj/ +| #L #L2 #_ #H * elim (rex_inv_bind_void … H) -H /3 width=3 by rexs_step_dx, conj/ +] +qed-. + +(* Advanced forward lemmas **************************************************) + +lemma rexs_fwd_bind_dx_void: ∀R. c_reflexive … R → + ∀p,I,L1,L2,V,T. L1 ⪤*[R, ⓑ{p,I}V.T] L2 → + L1.ⓧ ⪤*[R, T] L2.ⓧ. +#R #HR #p #I #L1 #L2 #V #T #H elim (rexs_inv_bind_void … H) -H // +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs_length.ma b/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs_length.ma new file mode 100644 index 000000000..9b1e1263b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs_length.ma @@ -0,0 +1,29 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rex_length.ma". +include "basic_2/i_static/rexs.ma". + +(* ITERATED EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ***) + +(* Forward lemmas with length for local environments ************************) + +(* Basic_2A1: uses: TC_lpx_sn_fwd_length *) +lemma rexs_fwd_length: ∀R,L1,L2,T. L1 ⪤*[R, T] L2 → |L1| = |L2|. +#R #L1 #L2 #T #H elim H -L2 +[ #L2 #HL12 >(rex_fwd_length … HL12) -HL12 // +| #L #L2 #_ #HL2 #IHL1 + >IHL1 -L1 >(rex_fwd_length … HL2) -HL2 // +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs_lex.ma b/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs_lex.ma new file mode 100644 index 000000000..c46649efe --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs_lex.ma @@ -0,0 +1,58 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/lex_tc.ma". +include "basic_2/static/req_fqup.ma". +include "basic_2/static/req_fsle.ma". +include "basic_2/i_static/rexs_fqup.ma". + +(* ITERATED EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ***) + +(* Properties with generic extension of a context sensitive relation ********) + +lemma rexs_lex: ∀R. c_reflexive … R → + ∀L1,L2,T. L1 ⪤[CTC … R] L2 → L1 ⪤*[R, T] L2. +#R #HR #L1 #L2 #T * +/5 width=7 by rexs_tc, sex_inv_tc_dx, sex_co, ext2_inv_tc, ext2_refl/ +qed. + +lemma rexs_lex_req: ∀R. c_reflexive … R → + ∀L1,L. L1 ⪤[CTC … R] L → ∀L2,T. L ≡[T] L2 → + L1 ⪤*[R, T] L2. +/3 width=3 by rexs_lex, rexs_step_dx, req_fwd_rex/ qed. + +(* Inversion lemmas with generic extension of a context sensitive relation **) + +(* Note: s_rs_transitive_lex_inv_isid could be invoked in the last auto but makes it too slow *) +lemma rexs_inv_lex_req: ∀R. c_reflexive … R → + rex_fsge_compatible R → + s_rs_transitive … R (λ_.lex R) → + req_transitive R → + ∀L1,L2,T. L1 ⪤*[R, T] L2 → + ∃∃L. L1 ⪤[CTC … R] L & L ≡[T] L2. +#R #H1R #H2R #H3R #H4R #L1 #L2 #T #H +lapply (s_rs_transitive_lex_inv_isid … H3R) -H3R #H3R +@(rexs_ind_sn … H1R … H) -H -L2 +[ /4 width=3 by req_refl, lex_refl, inj, ex2_intro/ +| #L0 #L2 #_ #HL02 * #L * #f0 #Hf0 #HL1 #HL0 + lapply (req_rex_trans … HL0 … HL02) -L0 // * #f1 #Hf1 #HL2 + elim (sex_sdj_split … ceq_ext … HL2 f0 ?) -HL2 + [ #L0 #HL0 #HL02 |*: /2 width=1 by ext2_refl, sdj_isid_dx/ ] + lapply (sex_sdj … HL0 f1 ?) /2 width=1 by sdj_isid_sn/ #H + elim (frees_sex_conf … Hf1 … H) // -H2R -H #f2 #Hf2 #Hf21 + lapply (sle_sex_trans … HL02 … Hf21) -f1 // #HL02 + lapply (sex_co ?? cfull (CTC … (cext2 R)) … HL1) -HL1 /2 width=1 by ext2_inv_tc/ #HL1 + /8 width=11 by sex_inv_tc_dx, sex_tc_dx, sex_co, ext2_tc, ext2_refl, step, ex2_intro/ (**) (* full auto too slow *) +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs_rexs.ma b/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs_rexs.ma new file mode 100644 index 000000000..7b0b83d48 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/i_static/rexs_rexs.ma @@ -0,0 +1,33 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rex_fsle.ma". +include "basic_2/i_static/rexs.ma". + +(* ITERATED EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ***) + +(* Advanced properties ******************************************************) + +lemma rexs_sym: ∀R. rex_fsge_compatible R → + (∀L1,L2,T1,T2. R L1 T1 T2 → R L2 T2 T1) → + ∀T. symmetric … (rexs R T). +#R #H1R #H2R #T #L1 #L2 #H elim H -L2 +/4 width=3 by rex_sym, rexs_step_sn, inj/ +qed-. + +(* Main properties **********************************************************) + +theorem rexs_trans: ∀R,T. Transitive … (rexs R T). +#R #T #L1 #L #HL1 #L2 #HL2 @(trans_TC … HL1 HL2) (**) (* auto fails *) +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs.ma b/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs.ma deleted file mode 100644 index da4959b88..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs.ma +++ /dev/null @@ -1,195 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "ground_2/lib/star.ma". -include "basic_2/notation/relations/relationstarstar_4.ma". -include "basic_2/static/lfxs.ma". - -(* ITERATED EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ***) - -definition tc_lfxs (R): term → relation lenv ≝ CTC … (lfxs R). - -interpretation "iterated extension on referred entries (local environment)" - 'RelationStarStar R T L1 L2 = (tc_lfxs R T L1 L2). - -(* Basic properties *********************************************************) - -lemma tc_lfxs_step_dx: ∀R,L1,L,T. L1 ⪤**[R, T] L → - ∀L2. L ⪤*[R, T] L2 → L1 ⪤**[R, T] L2. -#R #L1 #L2 #T #HL1 #L2 @step @HL1 (**) (* auto fails *) -qed-. - -lemma tc_lfxs_step_sn: ∀R,L1,L,T. L1 ⪤*[R, T] L → - ∀L2. L ⪤**[R, T] L2 → L1 ⪤**[R, T] L2. -#R #L1 #L2 #T #HL1 #L2 @TC_strap @HL1 (**) (* auto fails *) -qed-. - -lemma tc_lfxs_atom: ∀R,I. ⋆ ⪤**[R, ⓪{I}] ⋆. -/2 width=1 by inj/ qed. - -lemma tc_lfxs_sort: ∀R,I,L1,L2,V1,V2,s. - L1 ⪤**[R, ⋆s] L2 → L1.ⓑ{I}V1 ⪤**[R, ⋆s] L2.ⓑ{I}V2. -#R #I #L1 #L2 #V1 #V2 #s #H elim H -L2 -/3 width=4 by lfxs_sort, tc_lfxs_step_dx, inj/ -qed. - -lemma tc_lfxs_pair: ∀R. (∀L. reflexive … (R L)) → - ∀I,L1,L2,V. L1 ⪤**[R, V] L2 → - L1.ⓑ{I}V ⪤**[R, #0] L2.ⓑ{I}V. -#R #HR #I #L1 #L2 #V #H elim H -L2 -/3 width=5 by lfxs_pair, tc_lfxs_step_dx, inj/ -qed. - -lemma tc_lfxs_unit: ∀R,f,I,L1,L2. 𝐈⦃f⦄ → L1 ⪤*[cext2 R, cfull, f] L2 → - L1.ⓤ{I} ⪤**[R, #0] L2.ⓤ{I}. -/3 width=3 by lfxs_unit, inj/ qed. - -lemma tc_lfxs_lref: ∀R,I,L1,L2,V1,V2,i. - L1 ⪤**[R, #i] L2 → L1.ⓑ{I}V1 ⪤**[R, #↑i] L2.ⓑ{I}V2. -#R #I #L1 #L2 #V1 #V2 #i #H elim H -L2 -/3 width=4 by lfxs_lref, tc_lfxs_step_dx, inj/ -qed. - -lemma tc_lfxs_gref: ∀R,I,L1,L2,V1,V2,l. - L1 ⪤**[R, §l] L2 → L1.ⓑ{I}V1 ⪤**[R, §l] L2.ⓑ{I}V2. -#R #I #L1 #L2 #V1 #V2 #l #H elim H -L2 -/3 width=4 by lfxs_gref, tc_lfxs_step_dx, inj/ -qed. - -lemma tc_lfxs_co: ∀R1,R2. (∀L,T1,T2. R1 L T1 T2 → R2 L T1 T2) → - ∀L1,L2,T. L1 ⪤**[R1, T] L2 → L1 ⪤**[R2, T] L2. -#R1 #R2 #HR #L1 #L2 #T #H elim H -L2 -/4 width=5 by lfxs_co, tc_lfxs_step_dx, inj/ -qed-. - -(* Basic inversion lemmas ***************************************************) - -(* Basic_2A1: uses: TC_lpx_sn_inv_atom1 *) -lemma tc_lfxs_inv_atom_sn: ∀R,I,Y2. ⋆ ⪤**[R, ⓪{I}] Y2 → Y2 = ⋆. -#R #I #Y2 #H elim H -Y2 /3 width=3 by inj, lfxs_inv_atom_sn/ -qed-. - -(* Basic_2A1: uses: TC_lpx_sn_inv_atom2 *) -lemma tc_lfxs_inv_atom_dx: ∀R,I,Y1. Y1 ⪤**[R, ⓪{I}] ⋆ → Y1 = ⋆. -#R #I #Y1 #H @(TC_ind_dx ??????? H) -Y1 -/3 width=3 by inj, lfxs_inv_atom_dx/ -qed-. - -lemma tc_lfxs_inv_sort: ∀R,Y1,Y2,s. Y1 ⪤**[R, ⋆s] Y2 → - ∨∨ Y1 = ⋆ ∧ Y2 = ⋆ - | ∃∃I1,I2,L1,L2. L1 ⪤**[R, ⋆s] L2 & - Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. -#R #Y1 #Y2 #s #H elim H -Y2 -[ #Y2 #H elim (lfxs_inv_sort … H) -H * - /4 width=8 by ex3_4_intro, inj, or_introl, or_intror, conj/ -| #Y #Y2 #_ #H elim (lfxs_inv_sort … H) -H * - [ #H #H2 * * /3 width=7 by ex3_4_intro, or_introl, or_intror, conj/ - | #I #I2 #L #L2 #HL2 #H #H2 * * - [ #H1 #H0 destruct - | #I1 #I0 #L1 #L0 #HL10 #H1 #H0 destruct - /4 width=7 by ex3_4_intro, tc_lfxs_step_dx, or_intror/ - ] - ] -] -qed-. - -lemma tc_lfxs_inv_gref: ∀R,Y1,Y2,l. Y1 ⪤**[R, §l] Y2 → - ∨∨ Y1 = ⋆ ∧ Y2 = ⋆ - | ∃∃I1,I2,L1,L2. L1 ⪤**[R, §l] L2 & - Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. -#R #Y1 #Y2 #l #H elim H -Y2 -[ #Y2 #H elim (lfxs_inv_gref … H) -H * - /4 width=8 by ex3_4_intro, inj, or_introl, or_intror, conj/ -| #Y #Y2 #_ #H elim (lfxs_inv_gref … H) -H * - [ #H #H2 * * /3 width=7 by ex3_4_intro, or_introl, or_intror, conj/ - | #I #I2 #L #L2 #HL2 #H #H2 * * - [ #H1 #H0 destruct - | #I1 #I0 #L1 #L0 #HL10 #H1 #H0 destruct - /4 width=7 by ex3_4_intro, tc_lfxs_step_dx, or_intror/ - ] - ] -] -qed-. - -lemma tc_lfxs_inv_bind: ∀R. (∀L. reflexive … (R L)) → - ∀p,I,L1,L2,V,T. L1 ⪤**[R, ⓑ{p,I}V.T] L2 → - L1 ⪤**[R, V] L2 ∧ L1.ⓑ{I}V ⪤**[R, T] L2.ⓑ{I}V. -#R #HR #p #I #L1 #L2 #V #T #H elim H -L2 -[ #L2 #H elim (lfxs_inv_bind … V ? H) -H /3 width=1 by inj, conj/ -| #L #L2 #_ #H * elim (lfxs_inv_bind … V ? H) -H /3 width=3 by tc_lfxs_step_dx, conj/ -] -qed-. - -lemma tc_lfxs_inv_flat: ∀R,I,L1,L2,V,T. L1 ⪤**[R, ⓕ{I}V.T] L2 → - L1 ⪤**[R, V] L2 ∧ L1 ⪤**[R, T] L2. -#R #I #L1 #L2 #V #T #H elim H -L2 -[ #L2 #H elim (lfxs_inv_flat … H) -H /3 width=1 by inj, conj/ -| #L #L2 #_ #H * elim (lfxs_inv_flat … H) -H /3 width=3 by tc_lfxs_step_dx, conj/ -] -qed-. - -(* Advanced inversion lemmas ************************************************) - -lemma tc_lfxs_inv_sort_bind_sn: ∀R,I1,Y2,L1,s. L1.ⓘ{I1} ⪤**[R, ⋆s] Y2 → - ∃∃I2,L2. L1 ⪤**[R, ⋆s] L2 & Y2 = L2.ⓘ{I2}. -#R #I1 #Y2 #L1 #s #H elim (tc_lfxs_inv_sort … H) -H * -[ #H destruct -| #Z #I2 #Y1 #L2 #Hs #H1 #H2 destruct /2 width=4 by ex2_2_intro/ -] -qed-. - -lemma tc_lfxs_inv_sort_bind_dx: ∀R,I2,Y1,L2,s. Y1 ⪤**[R, ⋆s] L2.ⓘ{I2} → - ∃∃I1,L1. L1 ⪤**[R, ⋆s] L2 & Y1 = L1.ⓘ{I1}. -#R #I2 #Y1 #L2 #s #H elim (tc_lfxs_inv_sort … H) -H * -[ #_ #H destruct -| #I1 #Z #L1 #Y2 #Hs #H1 #H2 destruct /2 width=4 by ex2_2_intro/ -] -qed-. - -lemma tc_lfxs_inv_gref_bind_sn: ∀R,I1,Y2,L1,l. L1.ⓘ{I1} ⪤**[R, §l] Y2 → - ∃∃I2,L2. L1 ⪤**[R, §l] L2 & Y2 = L2.ⓘ{I2}. -#R #I1 #Y2 #L1 #l #H elim (tc_lfxs_inv_gref … H) -H * -[ #H destruct -| #Z #I2 #Y1 #L2 #Hl #H1 #H2 destruct /2 width=4 by ex2_2_intro/ -] -qed-. - -lemma tc_lfxs_inv_gref_bind_dx: ∀R,I2,Y1,L2,l. Y1 ⪤**[R, §l] L2.ⓘ{I2} → - ∃∃I1,L1. L1 ⪤**[R, §l] L2 & Y1 = L1.ⓘ{I1}. -#R #I2 #Y1 #L2 #l #H elim (tc_lfxs_inv_gref … H) -H * -[ #_ #H destruct -| #I1 #Z #L1 #Y2 #Hl #H1 #H2 destruct /2 width=4 by ex2_2_intro/ -] -qed-. - -(* Basic forward lemmas *****************************************************) - -lemma tc_lfxs_fwd_pair_sn: ∀R,I,L1,L2,V,T. L1 ⪤**[R, ②{I}V.T] L2 → L1 ⪤**[R, V] L2. -#R #I #L1 #L2 #V #T #H elim H -L2 -/3 width=5 by lfxs_fwd_pair_sn, tc_lfxs_step_dx, inj/ -qed-. - -lemma tc_lfxs_fwd_bind_dx: ∀R. (∀L. reflexive … (R L)) → - ∀p,I,L1,L2,V,T. L1 ⪤**[R, ⓑ{p,I}V.T] L2 → - L1.ⓑ{I}V ⪤**[R, T] L2.ⓑ{I}V. -#R #HR #p #I #L1 #L2 #V #T #H elim (tc_lfxs_inv_bind … H) -H // -qed-. - -lemma tc_lfxs_fwd_flat_dx: ∀R,I,L1,L2,V,T. L1 ⪤**[R, ⓕ{I}V.T] L2 → L1 ⪤**[R, T] L2. -#R #I #L1 #L2 #V #T #H elim (tc_lfxs_inv_flat … H) -H // -qed-. - -(* Basic_2A1: removed theorems 2: - TC_lpx_sn_inv_pair1 TC_lpx_sn_inv_pair2 -*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_drops.ma deleted file mode 100644 index 006ca8f1a..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_drops.ma +++ /dev/null @@ -1,72 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/relocation/lreq_lreq.ma". -include "basic_2/static/lfxs_drops.ma". -include "basic_2/i_static/tc_lfxs.ma". - -(* ITERATED EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ***) - -definition tc_f_dedropable_sn: predicate (relation3 lenv term term) ≝ - λR. ∀b,f,L1,K1. ⬇*[b, f] L1 ≘ K1 → - ∀K2,T. K1 ⪤**[R, T] K2 → ∀U. ⬆*[f] T ≘ U → - ∃∃L2. L1 ⪤**[R, U] L2 & ⬇*[b, f] L2 ≘ K2 & L1 ≡[f] L2. - -definition tc_f_dropable_sn: predicate (relation3 lenv term term) ≝ - λR. ∀b,f,L1,K1. ⬇*[b, f] L1 ≘ K1 → 𝐔⦃f⦄ → - ∀L2,U. L1 ⪤**[R, U] L2 → ∀T. ⬆*[f] T ≘ U → - ∃∃K2. K1 ⪤**[R, T] K2 & ⬇*[b, f] L2 ≘ K2. - -definition tc_f_dropable_dx: predicate (relation3 lenv term term) ≝ - λR. ∀L1,L2,U. L1 ⪤**[R, U] L2 → - ∀b,f,K2. ⬇*[b, f] L2 ≘ K2 → 𝐔⦃f⦄ → ∀T. ⬆*[f] T ≘ U → - ∃∃K1. ⬇*[b, f] L1 ≘ K1 & K1 ⪤**[R, T] K2. - -(* Properties with generic slicing for local environments *******************) - -lemma dedropable_sn_CTC: ∀R. f_dedropable_sn R → tc_f_dedropable_sn R. -#R #HR #b #f #L1 #K1 #HLK1 #K2 #T #H elim H -K2 -[ #K2 #HK12 #U #HTU elim (HR … HLK1 … HK12 … HTU) -K1 -T -HR - /3 width=4 by ex3_intro, inj/ -| #K #K2 #_ #HK2 #IH #U #HTU -HLK1 - elim (IH … HTU) -IH #L #HL1 #HLK - elim (HR … HLK … HK2 … HTU) -K -T -HR - /3 width=6 by lreq_trans, tc_lfxs_step_dx, ex3_intro/ -] -qed-. - -(* Inversion lemmas with generic slicing for local environments *************) - -lemma dropable_sn_CTC: ∀R. f_dropable_sn R → tc_f_dropable_sn R. -#R #HR #b #f #L1 #K1 #HLK1 #Hf #L2 #U #H elim H -L2 -[ #L2 #HL12 #T #HTU elim (HR … HLK1 … HL12 … HTU) -L1 -U -HR - /3 width=3 by inj, ex2_intro/ -| #L #L2 #_ #HL2 #IH #T #HTU -HLK1 - elim (IH … HTU) -IH #K #HK1 #HLK - elim (HR … HLK … HL2 … HTU) -L -U -HR - /3 width=3 by tc_lfxs_step_dx, ex2_intro/ -] -qed-. - -lemma dropable_dx_CTC: ∀R. f_dropable_dx R → tc_f_dropable_dx R. -#R #HR #L1 #L2 #U #H elim H -L2 -[ #L2 #HL12 #b #f #K2 #HLK2 #Hf #T #HTU - elim (HR … HL12 … HLK2 … HTU) -L2 -U -HR - /3 width=3 by inj, ex2_intro/ -| #L #L2 #_ #HL2 #IH #b #f #K2 #HLK2 #Hf #T #HTU - elim (HR … HL2 … HLK2 … HTU) -L2 -HR // #K #HLK #HK2 - elim (IH … HLK … HTU) -IH -L -U - /3 width=5 by tc_lfxs_step_dx, ex2_intro/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_fqup.ma b/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_fqup.ma deleted file mode 100644 index 3392f1e1a..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_fqup.ma +++ /dev/null @@ -1,75 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfxs_fqup.ma". -include "basic_2/i_static/tc_lfxs.ma". - -(* ITERATED EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ***) - -(* Advanced properties ******************************************************) - -lemma tc_lfxs_refl: ∀R. c_reflexive … R → - ∀T. reflexive … (tc_lfxs R T). -/3 width=1 by lfxs_refl, inj/ qed. - -(* Basic_2A1: uses: TC_lpx_sn_pair TC_lpx_sn_pair_refl *) -lemma tc_lfxs_pair_refl: ∀R. c_reflexive … R → - ∀L,V1,V2. CTC … R L V1 V2 → ∀I,T. L.ⓑ{I}V1 ⪤**[R, T] L.ⓑ{I}V2. -#R #HR #L #V1 #V2 #H elim H -V2 -/3 width=3 by tc_lfxs_step_dx, lfxs_pair_refl, inj/ -qed. - -lemma tc_lfxs_tc: ∀R,L1,L2,T,f. 𝐈⦃f⦄ → TC … (lexs cfull (cext2 R) f) L1 L2 → - L1 ⪤**[R, T] L2. -#R #L1 #L2 #T #f #Hf #H elim H -L2 -[ elim (frees_total L1 T) | #L elim (frees_total L T) ] -/5 width=7 by lexs_sdj, tc_lfxs_step_dx, sdj_isid_sn, inj, ex2_intro/ -qed. - -(* Advanced eliminators *****************************************************) - -lemma tc_lfxs_ind_sn: ∀R. c_reflexive … R → - ∀L1,T. ∀Q:predicate …. Q L1 → - (∀L,L2. L1 ⪤**[R, T] L → L ⪤*[R, T] L2 → Q L → Q L2) → - ∀L2. L1 ⪤**[R, T] L2 → Q L2. -#R #HR #L1 #T #Q #HL1 #IHL1 #L2 #HL12 -@(TC_star_ind … HL1 IHL1 … HL12) /2 width=1 by lfxs_refl/ -qed-. - -lemma tc_lfxs_ind_dx: ∀R. c_reflexive … R → - ∀L2,T. ∀Q:predicate …. Q L2 → - (∀L1,L. L1 ⪤*[R, T] L → L ⪤**[R, T] L2 → Q L → Q L1) → - ∀L1. L1 ⪤**[R, T] L2 → Q L1. -#R #HR #L2 #Q #HL2 #IHL2 #L1 #HL12 -@(TC_star_ind_dx … HL2 IHL2 … HL12) /2 width=4 by lfxs_refl/ -qed-. - -(* Advanced inversion lemmas ************************************************) - -lemma tc_lfxs_inv_bind_void: ∀R. c_reflexive … R → - ∀p,I,L1,L2,V,T. L1 ⪤**[R, ⓑ{p,I}V.T] L2 → - L1 ⪤**[R, V] L2 ∧ L1.ⓧ ⪤**[R, T] L2.ⓧ. -#R #HR #p #I #L1 #L2 #V #T #H @(tc_lfxs_ind_sn … HR … H) -L2 -[ /3 width=1 by tc_lfxs_refl, conj/ -| #L #L2 #_ #H * elim (lfxs_inv_bind_void … H) -H /3 width=3 by tc_lfxs_step_dx, conj/ -] -qed-. - -(* Advanced forward lemmas **************************************************) - -lemma tc_lfxs_fwd_bind_dx_void: ∀R. c_reflexive … R → - ∀p,I,L1,L2,V,T. L1 ⪤**[R, ⓑ{p,I}V.T] L2 → - L1.ⓧ ⪤**[R, T] L2.ⓧ. -#R #HR #p #I #L1 #L2 #V #T #H elim (tc_lfxs_inv_bind_void … H) -H // -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_length.ma b/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_length.ma deleted file mode 100644 index 8280cd80c..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_length.ma +++ /dev/null @@ -1,29 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfxs_length.ma". -include "basic_2/i_static/tc_lfxs.ma". - -(* ITERATED EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ***) - -(* Forward lemmas with length for local environments ************************) - -(* Basic_2A1: uses: TC_lpx_sn_fwd_length *) -lemma tc_lfxs_fwd_length: ∀R,L1,L2,T. L1 ⪤**[R, T] L2 → |L1| = |L2|. -#R #L1 #L2 #T #H elim H -L2 -[ #L2 #HL12 >(lfxs_fwd_length … HL12) -HL12 // -| #L #L2 #_ #HL2 #IHL1 - >IHL1 -L1 >(lfxs_fwd_length … HL2) -HL2 // -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_lex.ma b/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_lex.ma deleted file mode 100644 index 918191c8f..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_lex.ma +++ /dev/null @@ -1,58 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/relocation/lex_tc.ma". -include "basic_2/static/lfeq_fqup.ma". -include "basic_2/static/lfeq_fsle.ma". -include "basic_2/i_static/tc_lfxs_fqup.ma". - -(* ITERATED EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ***) - -(* Properties with generic extension of a context sensitive relation ********) - -lemma tc_lfxs_lex: ∀R. c_reflexive … R → - ∀L1,L2,T. L1 ⪤[CTC … R] L2 → L1 ⪤**[R, T] L2. -#R #HR #L1 #L2 #T * -/5 width=7 by tc_lfxs_tc, lexs_inv_tc_dx, lexs_co, ext2_inv_tc, ext2_refl/ -qed. - -lemma tc_lfxs_lex_lfeq: ∀R. c_reflexive … R → - ∀L1,L. L1 ⪤[CTC … R] L → ∀L2,T. L ≡[T] L2 → - L1 ⪤**[R, T] L2. -/3 width=3 by tc_lfxs_lex, tc_lfxs_step_dx, lfeq_fwd_lfxs/ qed. - -(* Inversion lemmas with generic extension of a context sensitive relation **) - -(* Note: s_rs_transitive_lex_inv_isid could be invoked in the last auto but makes it too slow *) -lemma tc_lfxs_inv_lex_lfeq: ∀R. c_reflexive … R → - lfxs_fsge_compatible R → - s_rs_transitive … R (λ_.lex R) → - lfeq_transitive R → - ∀L1,L2,T. L1 ⪤**[R, T] L2 → - ∃∃L. L1 ⪤[CTC … R] L & L ≡[T] L2. -#R #H1R #H2R #H3R #H4R #L1 #L2 #T #H -lapply (s_rs_transitive_lex_inv_isid … H3R) -H3R #H3R -@(tc_lfxs_ind_sn … H1R … H) -H -L2 -[ /4 width=3 by lfeq_refl, lex_refl, inj, ex2_intro/ -| #L0 #L2 #_ #HL02 * #L * #f0 #Hf0 #HL1 #HL0 - lapply (lfeq_lfxs_trans … HL0 … HL02) -L0 // * #f1 #Hf1 #HL2 - elim (lexs_sdj_split … ceq_ext … HL2 f0 ?) -HL2 - [ #L0 #HL0 #HL02 |*: /2 width=1 by ext2_refl, sdj_isid_dx/ ] - lapply (lexs_sdj … HL0 f1 ?) /2 width=1 by sdj_isid_sn/ #H - elim (frees_lexs_conf … Hf1 … H) // -H2R -H #f2 #Hf2 #Hf21 - lapply (sle_lexs_trans … HL02 … Hf21) -f1 // #HL02 - lapply (lexs_co ?? cfull (CTC … (cext2 R)) … HL1) -HL1 /2 width=1 by ext2_inv_tc/ #HL1 - /8 width=11 by lexs_inv_tc_dx, lexs_tc_dx, lexs_co, ext2_tc, ext2_refl, step, ex2_intro/ (**) (* full auto too slow *) -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_tc_lfxs.ma b/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_tc_lfxs.ma deleted file mode 100644 index dbcc54a3a..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_tc_lfxs.ma +++ /dev/null @@ -1,33 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfxs_fsle.ma". -include "basic_2/i_static/tc_lfxs.ma". - -(* ITERATED EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ***) - -(* Advanced properties ******************************************************) - -lemma tc_lfxs_sym: ∀R. lfxs_fsge_compatible R → - (∀L1,L2,T1,T2. R L1 T1 T2 → R L2 T2 T1) → - ∀T. symmetric … (tc_lfxs R T). -#R #H1R #H2R #T #L1 #L2 #H elim H -L2 -/4 width=3 by lfxs_sym, tc_lfxs_step_sn, inj/ -qed-. - -(* Main properties **********************************************************) - -theorem tc_lfxs_trans: ∀R,T. Transitive … (tc_lfxs R T). -#R #T #L1 #L #HL1 #L2 #HL2 @(trans_TC … HL1 HL2) (**) (* auto fails *) -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/relation_4.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/relation_4.ma new file mode 100644 index 000000000..6d4e20cea --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/relation_4.ma @@ -0,0 +1,19 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) + +notation "hvbox( L1 ⪤[ break term 46 R, break term 46 T ] break term 46 L2 )" + non associative with precedence 45 + for @{ 'Relation $R $T $L1 $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/relation_5.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/relation_5.ma new file mode 100644 index 000000000..6ad4bb3e7 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/relation_5.ma @@ -0,0 +1,19 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) + +notation "hvbox( L1 ⪤[ break term 46 R1, break term 46 R2, break term 46 f ] break term 46 L2 )" + non associative with precedence 45 + for @{ 'Relation $R1 $R2 $f $L1 $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/relationstar_4.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/relationstar_4.ma index 227148db7..bc5939830 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/relationstar_4.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/relationstar_4.ma @@ -14,6 +14,6 @@ (* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) -notation "hvbox( L1 ⪤ * [ break term 46 R, break term 46 T ] break term 46 L2 )" +notation "hvbox( L1 ⪤*[ break term 46 R, break term 46 T ] break term 46 L2 )" non associative with precedence 45 for @{ 'RelationStar $R $T $L1 $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/relationstar_5.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/relationstar_5.ma deleted file mode 100644 index 2324f561f..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/relationstar_5.ma +++ /dev/null @@ -1,19 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) - -notation "hvbox( L1 ⪤ * [ break term 46 R1, break term 46 R2, break term 46 f ] break term 46 L2 )" - non associative with precedence 45 - for @{ 'RelationStar $R1 $R2 $f $L1 $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/relationstarstar_4.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/relationstarstar_4.ma deleted file mode 100644 index abec81497..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/relationstarstar_4.ma +++ /dev/null @@ -1,19 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) - -notation "hvbox( L1 ⪤ * * [ break term 46 R, break term 46 T ] break term 46 L2 )" - non associative with precedence 45 - for @{ 'RelationStarStar $R $T $L1 $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma index 0a03101fd..f982cb1a6 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma @@ -15,7 +15,7 @@ include "ground_2/relocation/rtmap_coafter.ma". include "basic_2/notation/relations/rdropstar_3.ma". include "basic_2/notation/relations/rdropstar_4.ma". -include "basic_2/relocation/lreq.ma". +include "basic_2/relocation/seq.ma". include "basic_2/relocation/lifts_bind.ma". (* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_ctc.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_ctc.ma index cc485711c..de5d1ce48 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_ctc.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_ctc.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "ground_2/lib/star.ma". -include "basic_2/relocation/lreq_lreq.ma". +include "basic_2/relocation/seq_seq.ma". (* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************) @@ -67,6 +67,6 @@ lemma co_dedropable_sn_CTC: ∀R. co_dedropable_sn R → co_dedropable_sn (CTC /3 width=4 by inj, ex3_intro/ | #K #K2 #_ #HK2 #IH #f2 #Hf elim (IH … Hf) -IH -K1 #L #H1L1 #HLK #H2L1 elim (HR … HLK … HK2 … Hf) -HR -f1 -K - /3 width=6 by lreq_trans, step, ex3_intro/ + /3 width=6 by seq_trans, step, ex3_intro/ ] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lex.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lex.ma index b1c4f4d91..793ed1810 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lex.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lex.ma @@ -14,7 +14,7 @@ include "basic_2/relocation/lex.ma". include "basic_2/relocation/drops_cext2.ma". -include "basic_2/relocation/drops_lexs.ma". +include "basic_2/relocation/drops_sex.ma". (* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************) @@ -36,7 +36,7 @@ definition dropable_dx: predicate … ≝ lemma lex_liftable_dedropable_sn (R): c_reflexive … R → d_liftable2_sn … lifts R → dedropable_sn R. #R #H1R #H2R #b #f #L1 #K1 #HLK1 #K2 * #f1 #Hf1 #HK12 -elim (lexs_liftable_co_dedropable_sn … HLK1 … HK12) -K1 +elim (sex_liftable_co_dedropable_sn … HLK1 … HK12) -K1 /3 width=6 by cext2_d_liftable2_sn, cfull_lift_sn, ext2_refl, coafter_isid_dx, ex3_intro, ex2_intro/ qed-. @@ -45,14 +45,14 @@ qed-. (* Basic_2A1: was: lpx_sn_deliftable_dropable *) lemma lex_dropable_sn (R): dropable_sn R. #R #b #f #L1 #K1 #HLK1 #H1f #L2 * #f2 #Hf2 #HL12 -elim (lexs_co_dropable_sn … HLK1 … HL12) -L1 +elim (sex_co_dropable_sn … HLK1 … HL12) -L1 /3 width=3 by coafter_isid_dx, ex2_intro/ qed-. (* Basic_2A1: was: lpx_sn_dropable *) lemma lex_dropable_dx (R): dropable_dx R. #R #L1 #L2 * #f2 #Hf2 #HL12 #b #f #K2 #HLK2 #Hf -elim (lexs_co_dropable_dx … HL12 … HLK2) -L2 +elim (sex_co_dropable_dx … HL12 … HLK2) -L2 /3 width=5 by coafter_isid_dx, ex2_intro/ qed-. @@ -61,7 +61,7 @@ lemma lex_drops_conf_pair (R): ∀L1,L2. L1 ⪤[R] L2 → ∀b,f,I,K1,V1. ⬇*[b, f] L1 ≘ K1.ⓑ{I}V1 → 𝐔⦃f⦄ → ∃∃K2,V2. ⬇*[b, f] L2 ≘ K2.ⓑ{I}V2 & K1 ⪤[R] K2 & R K1 V1 V2. #R #L1 #L2 * #f2 #Hf2 #HL12 #b #f #I #K1 #V1 #HLK1 #Hf -elim (lexs_drops_conf_push … HL12 … HLK1 Hf f2) -L1 -Hf +elim (sex_drops_conf_push … HL12 … HLK1 Hf f2) -L1 -Hf [ #Z2 #K2 #HLK2 #HK12 #H elim (ext2_inv_pair_sn … H) -H #V2 #HV12 #H destruct /3 width=5 by ex3_2_intro, ex2_intro/ @@ -74,7 +74,7 @@ lemma lex_drops_trans_pair (R): ∀L1,L2. L1 ⪤[R] L2 → ∀b,f,I,K2,V2. ⬇*[b, f] L2 ≘ K2.ⓑ{I}V2 → 𝐔⦃f⦄ → ∃∃K1,V1. ⬇*[b, f] L1 ≘ K1.ⓑ{I}V1 & K1 ⪤[R] K2 & R K1 V1 V2. #R #L1 #L2 * #f2 #Hf2 #HL12 #b #f #I #K2 #V2 #HLK2 #Hf -elim (lexs_drops_trans_push … HL12 … HLK2 Hf f2) -L2 -Hf +elim (sex_drops_trans_push … HL12 … HLK2 Hf f2) -L2 -Hf [ #Z1 #K1 #HLK1 #HK12 #H elim (ext2_inv_pair_dx … H) -H #V1 #HV12 #H destruct /3 width=5 by ex3_2_intro, ex2_intro/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lexs.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lexs.ma deleted file mode 100644 index 11b08f3e6..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lexs.ma +++ /dev/null @@ -1,187 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/relocation/lifts_lifts_bind.ma". -include "basic_2/relocation/drops.ma". - -(* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************) - -(* Properties with entrywise extension of context-sensitive relations *******) - -(**) (* changed after commit 13218 *) -lemma lexs_co_dropable_sn: ∀RN,RP. co_dropable_sn (lexs RN RP). -#RN #RP #b #f #L1 #K1 #H elim H -f -L1 -K1 -[ #f #Hf #_ #f2 #X #H #f1 #Hf2 >(lexs_inv_atom1 … H) -X - /4 width=3 by lexs_atom, drops_atom, ex2_intro/ -| #f #I1 #L1 #K1 #_ #IH #Hf #f2 #X #H #f1 #Hf2 - elim (coafter_inv_nxx … Hf2) -Hf2 [2,3: // ] #g2 #Hg2 #H2 destruct - elim (lexs_inv_push1 … H) -H #I2 #L2 #HL12 #HI12 #H destruct - elim (IH … HL12 … Hg2) -g2 - /3 width=3 by isuni_inv_next, drops_drop, ex2_intro/ -| #f #I1 #J1 #L1 #K1 #HLK #HJI1 #IH #Hf #f2 #X #H #f1 #Hf2 - lapply (isuni_inv_push … Hf ??) -Hf [3: |*: // ] #Hf - lapply (drops_fwd_isid … HLK … Hf) -HLK #H0 destruct - lapply (liftsb_fwd_isid … HJI1 … Hf) -HJI1 #H0 destruct - elim (coafter_inv_pxx … Hf2) -Hf2 [1,3:* |*: // ] #g1 #g2 #Hg2 #H1 #H2 destruct - [ elim (lexs_inv_push1 … H) | elim (lexs_inv_next1 … H) ] -H #I2 #L2 #HL12 #HI12 #H destruct - elim (IH … HL12 … Hg2) -g2 -IH /2 width=1 by isuni_isid/ #K2 #HK12 #HLK2 - lapply (drops_fwd_isid … HLK2 … Hf) -HLK2 #H0 destruct - /4 width=3 by drops_refl, lexs_next, lexs_push, isid_push, ex2_intro/ -] -qed-. - -lemma lexs_liftable_co_dedropable_bi: ∀RN,RP. d_liftable2_sn … liftsb RN → d_liftable2_sn … liftsb RP → - ∀f2,L1,L2. L1 ⪤*[cfull, RP, f2] L2 → ∀f1,K1,K2. K1 ⪤*[RN, RP, f1] K2 → - ∀b,f. ⬇*[b, f] L1 ≘ K1 → ⬇*[b, f] L2 ≘ K2 → - f ~⊚ f1 ≘ f2 → L1 ⪤*[RN, RP, f2] L2. -#RN #RP #HRN #HRP #f2 #L1 #L2 #H elim H -f2 -L1 -L2 // -#g2 #I1 #I2 #L1 #L2 #HL12 #HI12 #IH #f1 #Y1 #Y2 #HK12 #b #f #HY1 #HY2 #H -[ elim (coafter_inv_xxn … H) [ |*: // ] -H #g #g1 #Hg2 #H1 #H2 destruct - elim (drops_inv_skip1 … HY1) -HY1 #J1 #K1 #HLK1 #HJI1 #H destruct - elim (drops_inv_skip1 … HY2) -HY2 #J2 #K2 #HLK2 #HJI2 #H destruct - elim (lexs_inv_next … HK12) -HK12 #HK12 #HJ12 - elim (HRN … HJ12 … HLK1 … HJI1) -HJ12 -HJI1 #Z #Hz - >(liftsb_mono … Hz … HJI2) -Z /3 width=9 by lexs_next/ -| elim (coafter_inv_xxp … H) [1,2: |*: // ] -H * - [ #g #g1 #Hg2 #H1 #H2 destruct - elim (drops_inv_skip1 … HY1) -HY1 #J1 #K1 #HLK1 #HJI1 #H destruct - elim (drops_inv_skip1 … HY2) -HY2 #J2 #K2 #HLK2 #HJI2 #H destruct - elim (lexs_inv_push … HK12) -HK12 #HK12 #HJ12 - elim (HRP … HJ12 … HLK1 … HJI1) -HJ12 -HJI1 #Z #Hz - >(liftsb_mono … Hz … HJI2) -Z /3 width=9 by lexs_push/ - | #g #Hg2 #H destruct - lapply (drops_inv_drop1 … HY1) -HY1 #HLK1 - lapply (drops_inv_drop1 … HY2) -HY2 #HLK2 - /3 width=9 by lexs_push/ - ] -] -qed-. - -lemma lexs_liftable_co_dedropable_sn: ∀RN,RP. (∀L. reflexive … (RN L)) → (∀L. reflexive … (RP L)) → - d_liftable2_sn … liftsb RN → d_liftable2_sn … liftsb RP → - co_dedropable_sn (lexs RN RP). -#RN #RP #H1RN #H1RP #H2RN #H2RP #b #f #L1 #K1 #H elim H -f -L1 -K1 -[ #f #Hf #X #f1 #H #f2 #Hf2 >(lexs_inv_atom1 … H) -X - /4 width=4 by drops_atom, lexs_atom, ex3_intro/ -| #f #I1 #L1 #K1 #_ #IHLK1 #K2 #f1 #HK12 #f2 #Hf2 - elim (coafter_inv_nxx … Hf2) -Hf2 [2,3: // ] #g2 #Hg2 #H destruct - elim (IHLK1 … HK12 … Hg2) -K1 - /3 width=6 by drops_drop, lexs_next, lexs_push, ex3_intro/ -| #f #I1 #J1 #L1 #K1 #HLK1 #HJI1 #IHLK1 #X #f1 #H #f2 #Hf2 - elim (coafter_inv_pxx … Hf2) -Hf2 [1,3: * |*: // ] #g1 #g2 #Hg2 #H1 #H2 destruct - [ elim (lexs_inv_push1 … H) | elim (lexs_inv_next1 … H) ] -H #J2 #K2 #HK12 #HJ12 #H destruct - [ elim (H2RP … HJ12 … HLK1 … HJI1) | elim (H2RN … HJ12 … HLK1 … HJI1) ] -J1 - elim (IHLK1 … HK12 … Hg2) -K1 - /3 width=6 by drops_skip, lexs_next, lexs_push, ex3_intro/ -] -qed-. - -fact lexs_dropable_dx_aux: ∀RN,RP,b,f,L2,K2. ⬇*[b, f] L2 ≘ K2 → 𝐔⦃f⦄ → - ∀f2,L1. L1 ⪤*[RN, RP, f2] L2 → ∀f1. f ~⊚ f1 ≘ f2 → - ∃∃K1. ⬇*[b, f] L1 ≘ K1 & K1 ⪤*[RN, RP, f1] K2. -#RN #RP #b #f #L2 #K2 #H elim H -f -L2 -K2 -[ #f #Hf #_ #f2 #X #H #f1 #Hf2 lapply (lexs_inv_atom2 … H) -H - #H destruct /4 width=3 by lexs_atom, drops_atom, ex2_intro/ -| #f #I2 #L2 #K2 #_ #IH #Hf #f2 #X #HX #f1 #Hf2 - elim (coafter_inv_nxx … Hf2) -Hf2 [2,3: // ] #g2 #Hg2 #H destruct - elim (lexs_inv_push2 … HX) -HX #I1 #L1 #HL12 #HI12 #H destruct - elim (IH … HL12 … Hg2) -L2 -I2 -g2 - /3 width=3 by drops_drop, isuni_inv_next, ex2_intro/ -| #f #I2 #J2 #L2 #K2 #_ #HJI2 #IH #Hf #f2 #X #HX #f1 #Hf2 - elim (coafter_inv_pxx … Hf2) -Hf2 [1,3: * |*: // ] #g1 #g2 #Hg2 #H1 #H2 destruct - [ elim (lexs_inv_push2 … HX) | elim (lexs_inv_next2 … HX) ] -HX #I1 #L1 #HL12 #HI12 #H destruct - elim (IH … HL12 … Hg2) -L2 -g2 /2 width=3 by isuni_fwd_push/ #K1 #HLK1 #HK12 - lapply (isuni_inv_push … Hf ??) -Hf [3,6: |*: // ] #Hf - lapply (liftsb_fwd_isid … HJI2 … Hf) #H destruct -HJI2 - lapply (drops_fwd_isid … HLK1 … Hf) #H destruct -HLK1 - /4 width=5 by lexs_next, lexs_push, drops_refl, isid_push, ex2_intro/ -] -qed-. - -lemma lexs_co_dropable_dx: ∀RN,RP. co_dropable_dx (lexs RN RP). -/2 width=5 by lexs_dropable_dx_aux/ qed-. - -lemma lexs_drops_conf_next: ∀RN,RP. - ∀f2,L1,L2. L1 ⪤*[RN, RP, f2] L2 → - ∀b,f,I1,K1. ⬇*[b, f] L1 ≘ K1.ⓘ{I1} → 𝐔⦃f⦄ → - ∀f1. f ~⊚ ↑f1 ≘ f2 → - ∃∃I2,K2. ⬇*[b, f] L2 ≘ K2.ⓘ{I2} & K1 ⪤*[RN, RP, f1] K2 & RN K1 I1 I2. -#RN #RP #f2 #L1 #L2 #HL12 #b #f #I1 #K1 #HLK1 #Hf #f1 #Hf2 -elim (lexs_co_dropable_sn … HLK1 … Hf … HL12 … Hf2) -L1 -f2 -Hf -#X #HX #HLK2 elim (lexs_inv_next1 … HX) -HX -#I2 #K2 #HK12 #HI12 #H destruct /2 width=5 by ex3_2_intro/ -qed-. - -lemma lexs_drops_conf_push: ∀RN,RP. - ∀f2,L1,L2. L1 ⪤*[RN, RP, f2] L2 → - ∀b,f,I1,K1. ⬇*[b, f] L1 ≘ K1.ⓘ{I1} → 𝐔⦃f⦄ → - ∀f1. f ~⊚ ⫯f1 ≘ f2 → - ∃∃I2,K2. ⬇*[b, f] L2 ≘ K2.ⓘ{I2} & K1 ⪤*[RN, RP, f1] K2 & RP K1 I1 I2. -#RN #RP #f2 #L1 #L2 #HL12 #b #f #I1 #K1 #HLK1 #Hf #f1 #Hf2 -elim (lexs_co_dropable_sn … HLK1 … Hf … HL12 … Hf2) -L1 -f2 -Hf -#X #HX #HLK2 elim (lexs_inv_push1 … HX) -HX -#I2 #K2 #HK12 #HI12 #H destruct /2 width=5 by ex3_2_intro/ -qed-. - -lemma lexs_drops_trans_next: ∀RN,RP,f2,L1,L2. L1 ⪤*[RN, RP, f2] L2 → - ∀b,f,I2,K2. ⬇*[b, f] L2 ≘ K2.ⓘ{I2} → 𝐔⦃f⦄ → - ∀f1. f ~⊚ ↑f1 ≘ f2 → - ∃∃I1,K1. ⬇*[b, f] L1 ≘ K1.ⓘ{I1} & K1 ⪤*[RN, RP, f1] K2 & RN K1 I1 I2. -#RN #RP #f2 #L1 #L2 #HL12 #b #f #I2 #K2 #HLK2 #Hf #f1 #Hf2 -elim (lexs_co_dropable_dx … HL12 … HLK2 … Hf … Hf2) -L2 -f2 -Hf -#X #HLK1 #HX elim (lexs_inv_next2 … HX) -HX -#I1 #K1 #HK12 #HI12 #H destruct /2 width=5 by ex3_2_intro/ -qed-. - -lemma lexs_drops_trans_push: ∀RN,RP,f2,L1,L2. L1 ⪤*[RN, RP, f2] L2 → - ∀b,f,I2,K2. ⬇*[b, f] L2 ≘ K2.ⓘ{I2} → 𝐔⦃f⦄ → - ∀f1. f ~⊚ ⫯f1 ≘ f2 → - ∃∃I1,K1. ⬇*[b, f] L1 ≘ K1.ⓘ{I1} & K1 ⪤*[RN, RP, f1] K2 & RP K1 I1 I2. -#RN #RP #f2 #L1 #L2 #HL12 #b #f #I2 #K2 #HLK2 #Hf #f1 #Hf2 -elim (lexs_co_dropable_dx … HL12 … HLK2 … Hf … Hf2) -L2 -f2 -Hf -#X #HLK1 #HX elim (lexs_inv_push2 … HX) -HX -#I1 #K1 #HK12 #HI12 #H destruct /2 width=5 by ex3_2_intro/ -qed-. - -lemma drops_lexs_trans_next: ∀RN,RP. (∀L. reflexive ? (RN L)) → (∀L. reflexive ? (RP L)) → - d_liftable2_sn … liftsb RN → d_liftable2_sn … liftsb RP → - ∀f1,K1,K2. K1 ⪤*[RN, RP, f1] K2 → - ∀b,f,I1,L1. ⬇*[b, f] L1.ⓘ{I1} ≘ K1 → - ∀f2. f ~⊚ f1 ≘ ↑f2 → - ∃∃I2,L2. ⬇*[b, f] L2.ⓘ{I2} ≘ K2 & L1 ⪤*[RN, RP, f2] L2 & RN L1 I1 I2 & L1.ⓘ{I1} ≡[f] L2.ⓘ{I2}. -#RN #RP #H1RN #H1RP #H2RN #H2RP #f1 #K1 #K2 #HK12 #b #f #I1 #L1 #HLK1 #f2 #Hf2 -elim (lexs_liftable_co_dedropable_sn … H1RN H1RP H2RN H2RP … HLK1 … HK12 … Hf2) -K1 -f1 -H1RN -H1RP -H2RN -H2RP -#X #HX #HLK2 #H1L12 elim (lexs_inv_next1 … HX) -HX -#I2 #L2 #H2L12 #HI12 #H destruct /2 width=6 by ex4_2_intro/ -qed-. - -lemma drops_lexs_trans_push: ∀RN,RP. (∀L. reflexive ? (RN L)) → (∀L. reflexive ? (RP L)) → - d_liftable2_sn … liftsb RN → d_liftable2_sn … liftsb RP → - ∀f1,K1,K2. K1 ⪤*[RN, RP, f1] K2 → - ∀b,f,I1,L1. ⬇*[b, f] L1.ⓘ{I1} ≘ K1 → - ∀f2. f ~⊚ f1 ≘ ⫯f2 → - ∃∃I2,L2. ⬇*[b, f] L2.ⓘ{I2} ≘ K2 & L1 ⪤*[RN, RP, f2] L2 & RP L1 I1 I2 & L1.ⓘ{I1} ≡[f] L2.ⓘ{I2}. -#RN #RP #H1RN #H1RP #H2RN #H2RP #f1 #K1 #K2 #HK12 #b #f #I1 #L1 #HLK1 #f2 #Hf2 -elim (lexs_liftable_co_dedropable_sn … H1RN H1RP H2RN H2RP … HLK1 … HK12 … Hf2) -K1 -f1 -H1RN -H1RP -H2RN -H2RP -#X #HX #HLK2 #H1L12 elim (lexs_inv_push1 … HX) -HX -#I2 #L2 #H2L12 #HI12 #H destruct /2 width=6 by ex4_2_intro/ -qed-. - -lemma drops_atom2_lexs_conf: ∀RN,RP,b,f1,L1. ⬇*[b, f1] L1 ≘ ⋆ → 𝐔⦃f1⦄ → - ∀f,L2. L1 ⪤*[RN, RP, f] L2 → - ∀f2. f1 ~⊚ f2 ≘f → ⬇*[b, f1] L2 ≘ ⋆. -#RN #RP #b #f1 #L1 #H1 #Hf1 #f #L2 #H2 #f2 #H3 -elim (lexs_co_dropable_sn … H1 … H2 … H3) // -H1 -H2 -H3 -Hf1 -#L #H #HL2 lapply (lexs_inv_atom1 … H) -H // -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lreq.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lreq.ma deleted file mode 100644 index d07ea2210..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lreq.ma +++ /dev/null @@ -1,61 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/relocation/drops_lexs.ma". - -(* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************) - -(* Properties with ranged equivalence for local environments ****************) - -lemma lreq_co_dedropable_sn: co_dedropable_sn lreq. -@lexs_liftable_co_dedropable_sn -/2 width=6 by cfull_lift_sn, ceq_lift_sn/ qed-. - -lemma lreq_co_dropable_sn: co_dropable_sn lreq. -@lexs_co_dropable_sn qed-. - -lemma lreq_co_dropable_dx: co_dropable_dx lreq. -@lexs_co_dropable_dx qed-. - -(* Basic_2A1: includes: lreq_drop_trans_be *) -lemma lreq_drops_trans_next: ∀f2,L1,L2. L1 ≡[f2] L2 → - ∀b,f,I,K2. ⬇*[b, f] L2 ≘ K2.ⓘ{I} → 𝐔⦃f⦄ → - ∀f1. f ~⊚ ↑f1 ≘ f2 → - ∃∃K1. ⬇*[b, f] L1 ≘ K1.ⓘ{I} & K1 ≡[f1] K2. -#f2 #L1 #L2 #HL12 #b #f #I2 #K2 #HLK2 #Hf #f1 #Hf2 -elim (lexs_drops_trans_next … HL12 … HLK2 Hf … Hf2) -f2 -L2 -Hf -#I1 #K1 #HLK1 #HK12 #H <(ceq_ext_inv_eq … H) -I2 -/2 width=3 by ex2_intro/ -qed-. - -(* Basic_2A1: includes: lreq_drop_conf_be *) -lemma lreq_drops_conf_next: ∀f2,L1,L2. L1 ≡[f2] L2 → - ∀b,f,I,K1. ⬇*[b, f] L1 ≘ K1.ⓘ{I} → 𝐔⦃f⦄ → - ∀f1. f ~⊚ ↑f1 ≘ f2 → - ∃∃K2. ⬇*[b, f] L2 ≘ K2.ⓘ{I} & K1 ≡[f1] K2. -#f2 #L1 #L2 #HL12 #b #f #I1 #K1 #HLK1 #Hf #f1 #Hf2 -elim (lreq_drops_trans_next … (lreq_sym … HL12) … HLK1 … Hf2) // -f2 -L1 -Hf -/3 width=3 by lreq_sym, ex2_intro/ -qed-. - -lemma drops_lreq_trans_next: ∀f1,K1,K2. K1 ≡[f1] K2 → - ∀b,f,I,L1. ⬇*[b, f] L1.ⓘ{I} ≘ K1 → - ∀f2. f ~⊚ f1 ≘ ↑f2 → - ∃∃L2. ⬇*[b, f] L2.ⓘ{I} ≘ K2 & L1 ≡[f2] L2 & L1.ⓘ{I} ≡[f] L2.ⓘ{I}. -#f1 #K1 #K2 #HK12 #b #f #I1 #L1 #HLK1 #f2 #Hf2 -elim (drops_lexs_trans_next … HK12 … HLK1 … Hf2) -f1 -K1 -/2 width=6 by cfull_lift_sn, ceq_lift_sn/ -#I2 #L2 #HLK2 #HL12 #H >(ceq_ext_inv_eq … H) -I1 -/2 width=4 by ex3_intro/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_ltc.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_ltc.ma index 3f999a5b2..62ae1ecc6 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_ltc.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_ltc.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "ground_2/lib/ltc.ma". -include "basic_2/relocation/lreq_lreq.ma". +include "basic_2/relocation/seq_seq.ma". (* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************) @@ -90,6 +90,6 @@ lemma co_dedropable_sn_ltc: ∀A,f. associative … f → | #n1 #n2 #K #K2 #_ #IH #HK2 #f2 #Hf elim (IH … Hf) -K1 -IH #L #H1L1 #HLK #H2L1 elim (HR … HLK … HK2 … Hf) -f1 -K -HR - /3 width=6 by lreq_trans, ltc_dx, ex3_intro/ + /3 width=6 by seq_trans, ltc_dx, ex3_intro/ ] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_seq.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_seq.ma new file mode 100644 index 000000000..a722d46e7 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_seq.ma @@ -0,0 +1,61 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/drops_sex.ma". + +(* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************) + +(* Properties with syntactic equivalence for selected local environments ****) + +lemma seq_co_dedropable_sn: co_dedropable_sn seq. +@sex_liftable_co_dedropable_sn +/2 width=6 by cfull_lift_sn, ceq_lift_sn/ qed-. + +lemma seq_co_dropable_sn: co_dropable_sn seq. +@sex_co_dropable_sn qed-. + +lemma seq_co_dropable_dx: co_dropable_dx seq. +@sex_co_dropable_dx qed-. + +(* Basic_2A1: includes: lreq_drop_trans_be *) +lemma seq_drops_trans_next: ∀f2,L1,L2. L1 ≡[f2] L2 → + ∀b,f,I,K2. ⬇*[b, f] L2 ≘ K2.ⓘ{I} → 𝐔⦃f⦄ → + ∀f1. f ~⊚ ↑f1 ≘ f2 → + ∃∃K1. ⬇*[b, f] L1 ≘ K1.ⓘ{I} & K1 ≡[f1] K2. +#f2 #L1 #L2 #HL12 #b #f #I2 #K2 #HLK2 #Hf #f1 #Hf2 +elim (sex_drops_trans_next … HL12 … HLK2 Hf … Hf2) -f2 -L2 -Hf +#I1 #K1 #HLK1 #HK12 #H <(ceq_ext_inv_eq … H) -I2 +/2 width=3 by ex2_intro/ +qed-. + +(* Basic_2A1: includes: lreq_drop_conf_be *) +lemma seq_drops_conf_next: ∀f2,L1,L2. L1 ≡[f2] L2 → + ∀b,f,I,K1. ⬇*[b, f] L1 ≘ K1.ⓘ{I} → 𝐔⦃f⦄ → + ∀f1. f ~⊚ ↑f1 ≘ f2 → + ∃∃K2. ⬇*[b, f] L2 ≘ K2.ⓘ{I} & K1 ≡[f1] K2. +#f2 #L1 #L2 #HL12 #b #f #I1 #K1 #HLK1 #Hf #f1 #Hf2 +elim (seq_drops_trans_next … (seq_sym … HL12) … HLK1 … Hf2) // -f2 -L1 -Hf +/3 width=3 by seq_sym, ex2_intro/ +qed-. + +lemma drops_seq_trans_next: ∀f1,K1,K2. K1 ≡[f1] K2 → + ∀b,f,I,L1. ⬇*[b, f] L1.ⓘ{I} ≘ K1 → + ∀f2. f ~⊚ f1 ≘ ↑f2 → + ∃∃L2. ⬇*[b, f] L2.ⓘ{I} ≘ K2 & L1 ≡[f2] L2 & L1.ⓘ{I} ≡[f] L2.ⓘ{I}. +#f1 #K1 #K2 #HK12 #b #f #I1 #L1 #HLK1 #f2 #Hf2 +elim (drops_sex_trans_next … HK12 … HLK1 … Hf2) -f1 -K1 +/2 width=6 by cfull_lift_sn, ceq_lift_sn/ +#I2 #L2 #HLK2 #HL12 #H >(ceq_ext_inv_eq … H) -I1 +/2 width=4 by ex3_intro/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_sex.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_sex.ma new file mode 100644 index 000000000..5f5dfdb3c --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_sex.ma @@ -0,0 +1,187 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/lifts_lifts_bind.ma". +include "basic_2/relocation/drops.ma". + +(* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************) + +(* Properties with entrywise extension of context-sensitive relations *******) + +(**) (* changed after commit 13218 *) +lemma sex_co_dropable_sn: ∀RN,RP. co_dropable_sn (sex RN RP). +#RN #RP #b #f #L1 #K1 #H elim H -f -L1 -K1 +[ #f #Hf #_ #f2 #X #H #f1 #Hf2 >(sex_inv_atom1 … H) -X + /4 width=3 by sex_atom, drops_atom, ex2_intro/ +| #f #I1 #L1 #K1 #_ #IH #Hf #f2 #X #H #f1 #Hf2 + elim (coafter_inv_nxx … Hf2) -Hf2 [2,3: // ] #g2 #Hg2 #H2 destruct + elim (sex_inv_push1 … H) -H #I2 #L2 #HL12 #HI12 #H destruct + elim (IH … HL12 … Hg2) -g2 + /3 width=3 by isuni_inv_next, drops_drop, ex2_intro/ +| #f #I1 #J1 #L1 #K1 #HLK #HJI1 #IH #Hf #f2 #X #H #f1 #Hf2 + lapply (isuni_inv_push … Hf ??) -Hf [3: |*: // ] #Hf + lapply (drops_fwd_isid … HLK … Hf) -HLK #H0 destruct + lapply (liftsb_fwd_isid … HJI1 … Hf) -HJI1 #H0 destruct + elim (coafter_inv_pxx … Hf2) -Hf2 [1,3:* |*: // ] #g1 #g2 #Hg2 #H1 #H2 destruct + [ elim (sex_inv_push1 … H) | elim (sex_inv_next1 … H) ] -H #I2 #L2 #HL12 #HI12 #H destruct + elim (IH … HL12 … Hg2) -g2 -IH /2 width=1 by isuni_isid/ #K2 #HK12 #HLK2 + lapply (drops_fwd_isid … HLK2 … Hf) -HLK2 #H0 destruct + /4 width=3 by drops_refl, sex_next, sex_push, isid_push, ex2_intro/ +] +qed-. + +lemma sex_liftable_co_dedropable_bi: ∀RN,RP. d_liftable2_sn … liftsb RN → d_liftable2_sn … liftsb RP → + ∀f2,L1,L2. L1 ⪤[cfull, RP, f2] L2 → ∀f1,K1,K2. K1 ⪤[RN, RP, f1] K2 → + ∀b,f. ⬇*[b, f] L1 ≘ K1 → ⬇*[b, f] L2 ≘ K2 → + f ~⊚ f1 ≘ f2 → L1 ⪤[RN, RP, f2] L2. +#RN #RP #HRN #HRP #f2 #L1 #L2 #H elim H -f2 -L1 -L2 // +#g2 #I1 #I2 #L1 #L2 #HL12 #HI12 #IH #f1 #Y1 #Y2 #HK12 #b #f #HY1 #HY2 #H +[ elim (coafter_inv_xxn … H) [ |*: // ] -H #g #g1 #Hg2 #H1 #H2 destruct + elim (drops_inv_skip1 … HY1) -HY1 #J1 #K1 #HLK1 #HJI1 #H destruct + elim (drops_inv_skip1 … HY2) -HY2 #J2 #K2 #HLK2 #HJI2 #H destruct + elim (sex_inv_next … HK12) -HK12 #HK12 #HJ12 + elim (HRN … HJ12 … HLK1 … HJI1) -HJ12 -HJI1 #Z #Hz + >(liftsb_mono … Hz … HJI2) -Z /3 width=9 by sex_next/ +| elim (coafter_inv_xxp … H) [1,2: |*: // ] -H * + [ #g #g1 #Hg2 #H1 #H2 destruct + elim (drops_inv_skip1 … HY1) -HY1 #J1 #K1 #HLK1 #HJI1 #H destruct + elim (drops_inv_skip1 … HY2) -HY2 #J2 #K2 #HLK2 #HJI2 #H destruct + elim (sex_inv_push … HK12) -HK12 #HK12 #HJ12 + elim (HRP … HJ12 … HLK1 … HJI1) -HJ12 -HJI1 #Z #Hz + >(liftsb_mono … Hz … HJI2) -Z /3 width=9 by sex_push/ + | #g #Hg2 #H destruct + lapply (drops_inv_drop1 … HY1) -HY1 #HLK1 + lapply (drops_inv_drop1 … HY2) -HY2 #HLK2 + /3 width=9 by sex_push/ + ] +] +qed-. + +lemma sex_liftable_co_dedropable_sn: ∀RN,RP. (∀L. reflexive … (RN L)) → (∀L. reflexive … (RP L)) → + d_liftable2_sn … liftsb RN → d_liftable2_sn … liftsb RP → + co_dedropable_sn (sex RN RP). +#RN #RP #H1RN #H1RP #H2RN #H2RP #b #f #L1 #K1 #H elim H -f -L1 -K1 +[ #f #Hf #X #f1 #H #f2 #Hf2 >(sex_inv_atom1 … H) -X + /4 width=4 by drops_atom, sex_atom, ex3_intro/ +| #f #I1 #L1 #K1 #_ #IHLK1 #K2 #f1 #HK12 #f2 #Hf2 + elim (coafter_inv_nxx … Hf2) -Hf2 [2,3: // ] #g2 #Hg2 #H destruct + elim (IHLK1 … HK12 … Hg2) -K1 + /3 width=6 by drops_drop, sex_next, sex_push, ex3_intro/ +| #f #I1 #J1 #L1 #K1 #HLK1 #HJI1 #IHLK1 #X #f1 #H #f2 #Hf2 + elim (coafter_inv_pxx … Hf2) -Hf2 [1,3: * |*: // ] #g1 #g2 #Hg2 #H1 #H2 destruct + [ elim (sex_inv_push1 … H) | elim (sex_inv_next1 … H) ] -H #J2 #K2 #HK12 #HJ12 #H destruct + [ elim (H2RP … HJ12 … HLK1 … HJI1) | elim (H2RN … HJ12 … HLK1 … HJI1) ] -J1 + elim (IHLK1 … HK12 … Hg2) -K1 + /3 width=6 by drops_skip, sex_next, sex_push, ex3_intro/ +] +qed-. + +fact sex_dropable_dx_aux: ∀RN,RP,b,f,L2,K2. ⬇*[b, f] L2 ≘ K2 → 𝐔⦃f⦄ → + ∀f2,L1. L1 ⪤[RN, RP, f2] L2 → ∀f1. f ~⊚ f1 ≘ f2 → + ∃∃K1. ⬇*[b, f] L1 ≘ K1 & K1 ⪤[RN, RP, f1] K2. +#RN #RP #b #f #L2 #K2 #H elim H -f -L2 -K2 +[ #f #Hf #_ #f2 #X #H #f1 #Hf2 lapply (sex_inv_atom2 … H) -H + #H destruct /4 width=3 by sex_atom, drops_atom, ex2_intro/ +| #f #I2 #L2 #K2 #_ #IH #Hf #f2 #X #HX #f1 #Hf2 + elim (coafter_inv_nxx … Hf2) -Hf2 [2,3: // ] #g2 #Hg2 #H destruct + elim (sex_inv_push2 … HX) -HX #I1 #L1 #HL12 #HI12 #H destruct + elim (IH … HL12 … Hg2) -L2 -I2 -g2 + /3 width=3 by drops_drop, isuni_inv_next, ex2_intro/ +| #f #I2 #J2 #L2 #K2 #_ #HJI2 #IH #Hf #f2 #X #HX #f1 #Hf2 + elim (coafter_inv_pxx … Hf2) -Hf2 [1,3: * |*: // ] #g1 #g2 #Hg2 #H1 #H2 destruct + [ elim (sex_inv_push2 … HX) | elim (sex_inv_next2 … HX) ] -HX #I1 #L1 #HL12 #HI12 #H destruct + elim (IH … HL12 … Hg2) -L2 -g2 /2 width=3 by isuni_fwd_push/ #K1 #HLK1 #HK12 + lapply (isuni_inv_push … Hf ??) -Hf [3,6: |*: // ] #Hf + lapply (liftsb_fwd_isid … HJI2 … Hf) #H destruct -HJI2 + lapply (drops_fwd_isid … HLK1 … Hf) #H destruct -HLK1 + /4 width=5 by sex_next, sex_push, drops_refl, isid_push, ex2_intro/ +] +qed-. + +lemma sex_co_dropable_dx: ∀RN,RP. co_dropable_dx (sex RN RP). +/2 width=5 by sex_dropable_dx_aux/ qed-. + +lemma sex_drops_conf_next: ∀RN,RP. + ∀f2,L1,L2. L1 ⪤[RN, RP, f2] L2 → + ∀b,f,I1,K1. ⬇*[b, f] L1 ≘ K1.ⓘ{I1} → 𝐔⦃f⦄ → + ∀f1. f ~⊚ ↑f1 ≘ f2 → + ∃∃I2,K2. ⬇*[b, f] L2 ≘ K2.ⓘ{I2} & K1 ⪤[RN, RP, f1] K2 & RN K1 I1 I2. +#RN #RP #f2 #L1 #L2 #HL12 #b #f #I1 #K1 #HLK1 #Hf #f1 #Hf2 +elim (sex_co_dropable_sn … HLK1 … Hf … HL12 … Hf2) -L1 -f2 -Hf +#X #HX #HLK2 elim (sex_inv_next1 … HX) -HX +#I2 #K2 #HK12 #HI12 #H destruct /2 width=5 by ex3_2_intro/ +qed-. + +lemma sex_drops_conf_push: ∀RN,RP. + ∀f2,L1,L2. L1 ⪤[RN, RP, f2] L2 → + ∀b,f,I1,K1. ⬇*[b, f] L1 ≘ K1.ⓘ{I1} → 𝐔⦃f⦄ → + ∀f1. f ~⊚ ⫯f1 ≘ f2 → + ∃∃I2,K2. ⬇*[b, f] L2 ≘ K2.ⓘ{I2} & K1 ⪤[RN, RP, f1] K2 & RP K1 I1 I2. +#RN #RP #f2 #L1 #L2 #HL12 #b #f #I1 #K1 #HLK1 #Hf #f1 #Hf2 +elim (sex_co_dropable_sn … HLK1 … Hf … HL12 … Hf2) -L1 -f2 -Hf +#X #HX #HLK2 elim (sex_inv_push1 … HX) -HX +#I2 #K2 #HK12 #HI12 #H destruct /2 width=5 by ex3_2_intro/ +qed-. + +lemma sex_drops_trans_next: ∀RN,RP,f2,L1,L2. L1 ⪤[RN, RP, f2] L2 → + ∀b,f,I2,K2. ⬇*[b, f] L2 ≘ K2.ⓘ{I2} → 𝐔⦃f⦄ → + ∀f1. f ~⊚ ↑f1 ≘ f2 → + ∃∃I1,K1. ⬇*[b, f] L1 ≘ K1.ⓘ{I1} & K1 ⪤[RN, RP, f1] K2 & RN K1 I1 I2. +#RN #RP #f2 #L1 #L2 #HL12 #b #f #I2 #K2 #HLK2 #Hf #f1 #Hf2 +elim (sex_co_dropable_dx … HL12 … HLK2 … Hf … Hf2) -L2 -f2 -Hf +#X #HLK1 #HX elim (sex_inv_next2 … HX) -HX +#I1 #K1 #HK12 #HI12 #H destruct /2 width=5 by ex3_2_intro/ +qed-. + +lemma sex_drops_trans_push: ∀RN,RP,f2,L1,L2. L1 ⪤[RN, RP, f2] L2 → + ∀b,f,I2,K2. ⬇*[b, f] L2 ≘ K2.ⓘ{I2} → 𝐔⦃f⦄ → + ∀f1. f ~⊚ ⫯f1 ≘ f2 → + ∃∃I1,K1. ⬇*[b, f] L1 ≘ K1.ⓘ{I1} & K1 ⪤[RN, RP, f1] K2 & RP K1 I1 I2. +#RN #RP #f2 #L1 #L2 #HL12 #b #f #I2 #K2 #HLK2 #Hf #f1 #Hf2 +elim (sex_co_dropable_dx … HL12 … HLK2 … Hf … Hf2) -L2 -f2 -Hf +#X #HLK1 #HX elim (sex_inv_push2 … HX) -HX +#I1 #K1 #HK12 #HI12 #H destruct /2 width=5 by ex3_2_intro/ +qed-. + +lemma drops_sex_trans_next: ∀RN,RP. (∀L. reflexive ? (RN L)) → (∀L. reflexive ? (RP L)) → + d_liftable2_sn … liftsb RN → d_liftable2_sn … liftsb RP → + ∀f1,K1,K2. K1 ⪤[RN, RP, f1] K2 → + ∀b,f,I1,L1. ⬇*[b, f] L1.ⓘ{I1} ≘ K1 → + ∀f2. f ~⊚ f1 ≘ ↑f2 → + ∃∃I2,L2. ⬇*[b, f] L2.ⓘ{I2} ≘ K2 & L1 ⪤[RN, RP, f2] L2 & RN L1 I1 I2 & L1.ⓘ{I1} ≡[f] L2.ⓘ{I2}. +#RN #RP #H1RN #H1RP #H2RN #H2RP #f1 #K1 #K2 #HK12 #b #f #I1 #L1 #HLK1 #f2 #Hf2 +elim (sex_liftable_co_dedropable_sn … H1RN H1RP H2RN H2RP … HLK1 … HK12 … Hf2) -K1 -f1 -H1RN -H1RP -H2RN -H2RP +#X #HX #HLK2 #H1L12 elim (sex_inv_next1 … HX) -HX +#I2 #L2 #H2L12 #HI12 #H destruct /2 width=6 by ex4_2_intro/ +qed-. + +lemma drops_sex_trans_push: ∀RN,RP. (∀L. reflexive ? (RN L)) → (∀L. reflexive ? (RP L)) → + d_liftable2_sn … liftsb RN → d_liftable2_sn … liftsb RP → + ∀f1,K1,K2. K1 ⪤[RN, RP, f1] K2 → + ∀b,f,I1,L1. ⬇*[b, f] L1.ⓘ{I1} ≘ K1 → + ∀f2. f ~⊚ f1 ≘ ⫯f2 → + ∃∃I2,L2. ⬇*[b, f] L2.ⓘ{I2} ≘ K2 & L1 ⪤[RN, RP, f2] L2 & RP L1 I1 I2 & L1.ⓘ{I1} ≡[f] L2.ⓘ{I2}. +#RN #RP #H1RN #H1RP #H2RN #H2RP #f1 #K1 #K2 #HK12 #b #f #I1 #L1 #HLK1 #f2 #Hf2 +elim (sex_liftable_co_dedropable_sn … H1RN H1RP H2RN H2RP … HLK1 … HK12 … Hf2) -K1 -f1 -H1RN -H1RP -H2RN -H2RP +#X #HX #HLK2 #H1L12 elim (sex_inv_push1 … HX) -HX +#I2 #L2 #H2L12 #HI12 #H destruct /2 width=6 by ex4_2_intro/ +qed-. + +lemma drops_atom2_sex_conf: ∀RN,RP,b,f1,L1. ⬇*[b, f1] L1 ≘ ⋆ → 𝐔⦃f1⦄ → + ∀f,L2. L1 ⪤[RN, RP, f] L2 → + ∀f2. f1 ~⊚ f2 ≘f → ⬇*[b, f1] L2 ≘ ⋆. +#RN #RP #b #f1 #L1 #H1 #Hf1 #f #L2 #H2 #f2 #H3 +elim (sex_co_dropable_sn … H1 … H2 … H3) // -H1 -H2 -H3 -Hf1 +#L #H #HL2 lapply (sex_inv_atom1 … H) -H // +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lex.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lex.ma index b1d925a3a..7e6aeae49 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lex.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lex.ma @@ -17,12 +17,12 @@ include "ground_2/pull/pull_4.ma". include "ground_2/relocation/rtmap_uni.ma". include "basic_2/notation/relations/relation_3.ma". include "basic_2/syntax/cext2.ma". -include "basic_2/relocation/lexs.ma". +include "basic_2/relocation/sex.ma". (* GENERIC EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS **************) definition lex (R): relation lenv ≝ - λL1,L2. ∃∃f. 𝐈⦃f⦄ & L1 ⪤*[cfull, cext2 R, f] L2. + λL1,L2. ∃∃f. 𝐈⦃f⦄ & L1 ⪤[cfull, cext2 R, f] L2. interpretation "generic extension (local environment)" 'Relation R L1 L2 = (lex R L1 L2). @@ -40,21 +40,21 @@ definition lex_transitive: relation (relation3 …) ≝ λR1,R2. (* Basic_2A1: was: lpx_sn_atom *) lemma lex_atom (R): ⋆ ⪤[R] ⋆. -/2 width=3 by lexs_atom, ex2_intro/ qed. +/2 width=3 by sex_atom, ex2_intro/ qed. lemma lex_bind (R): ∀I1,I2,K1,K2. K1 ⪤[R] K2 → cext2 R K1 I1 I2 → K1.ⓘ{I1} ⪤[R] K2.ⓘ{I2}. #R #I1 #I2 #K1 #K2 * #f #Hf #HK12 #HI12 -/3 width=3 by lexs_push, isid_push, ex2_intro/ +/3 width=3 by sex_push, isid_push, ex2_intro/ qed. (* Basic_2A1: was: lpx_sn_refl *) lemma lex_refl (R): c_reflexive … R → reflexive … (lex R). -/4 width=3 by lexs_refl, ext2_refl, ex2_intro/ qed. +/4 width=3 by sex_refl, ext2_refl, ex2_intro/ qed. lemma lex_co (R1) (R2): (∀L,T1,T2. R1 L T1 T2 → R2 L T1 T2) → ∀L1,L2. L1 ⪤[R1] L2 → L1 ⪤[R2] L2. -#R1 #R2 #HR #L1 #L2 * /5 width=7 by lexs_co, cext2_co, ex2_intro/ +#R1 #R2 #HR #L1 #L2 * /5 width=7 by sex_co, cext2_co, ex2_intro/ qed-. (* Advanced properties ******************************************************) @@ -75,27 +75,27 @@ lemma lex_pair (R): ∀I,K1,K2,V1,V2. K1 ⪤[R] K2 → R K1 V1 V2 → (* Basic_2A1: was: lpx_sn_inv_atom1: *) lemma lex_inv_atom_sn (R): ∀L2. ⋆ ⪤[R] L2 → L2 = ⋆. -#R #L2 * #f #Hf #H >(lexs_inv_atom1 … H) -L2 // +#R #L2 * #f #Hf #H >(sex_inv_atom1 … H) -L2 // qed-. lemma lex_inv_bind_sn (R): ∀I1,L2,K1. K1.ⓘ{I1} ⪤[R] L2 → ∃∃I2,K2. K1 ⪤[R] K2 & cext2 R K1 I1 I2 & L2 = K2.ⓘ{I2}. #R #I1 #L2 #K1 * #f #Hf #H -lapply (lexs_eq_repl_fwd … H (⫯f) ?) -H /2 width=1 by eq_push_inv_isid/ #H -elim (lexs_inv_push1 … H) -H #I2 #K2 #HK12 #HI12 #H destruct +lapply (sex_eq_repl_fwd … H (⫯f) ?) -H /2 width=1 by eq_push_inv_isid/ #H +elim (sex_inv_push1 … H) -H #I2 #K2 #HK12 #HI12 #H destruct /3 width=5 by ex2_intro, ex3_2_intro/ qed-. (* Basic_2A1: was: lpx_sn_inv_atom2 *) lemma lex_inv_atom_dx (R): ∀L1. L1 ⪤[R] ⋆ → L1 = ⋆. -#R #L1 * #f #Hf #H >(lexs_inv_atom2 … H) -L1 // +#R #L1 * #f #Hf #H >(sex_inv_atom2 … H) -L1 // qed-. lemma lex_inv_bind_dx (R): ∀I2,L1,K2. L1 ⪤[R] K2.ⓘ{I2} → ∃∃I1,K1. K1 ⪤[R] K2 & cext2 R K1 I1 I2 & L1 = K1.ⓘ{I1}. #R #I2 #L1 #K2 * #f #Hf #H -lapply (lexs_eq_repl_fwd … H (⫯f) ?) -H /2 width=1 by eq_push_inv_isid/ #H -elim (lexs_inv_push2 … H) -H #I1 #K1 #HK12 #HI12 #H destruct +lapply (sex_eq_repl_fwd … H (⫯f) ?) -H /2 width=1 by eq_push_inv_isid/ #H +elim (sex_inv_push2 … H) -H #I1 #K1 #HK12 #HI12 #H destruct /3 width=5 by ex3_2_intro, ex2_intro/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lex_length.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lex_length.ma index 9c87f6652..4f77163fe 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lex_length.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lex_length.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/relocation/lexs_length.ma". +include "basic_2/relocation/sex_length.ma". include "basic_2/relocation/lex.ma". (* GENERIC EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS **************) @@ -21,5 +21,5 @@ include "basic_2/relocation/lex.ma". (* Basic_2A1: was: lpx_sn_fwd_length *) lemma lex_fwd_length: ∀R,L1,L2. L1 ⪤[R] L2 → |L1| = |L2|. -#R #L1 #L2 * /2 width=4 by lexs_fwd_length/ +#R #L1 #L2 * /2 width=4 by sex_fwd_length/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lex_tc.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lex_tc.ma index a9ed87b92..00b9db820 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lex_tc.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lex_tc.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "basic_2/syntax/ext2_tc.ma". -include "basic_2/relocation/lexs_tc.ma". +include "basic_2/relocation/sex_tc.ma". include "basic_2/relocation/lex.ma". alias symbol "subseteq" = "relation inclusion". @@ -26,7 +26,7 @@ alias symbol "subseteq" = "relation inclusion". lemma lex_inv_CTC (R): c_reflexive … R → lex (CTC … R) ⊆ TC … (lex R). #R #HR #L1 #L2 * -/5 width=11 by lexs_inv_tc_dx, lexs_co, ext2_inv_tc, ext2_refl, monotonic_TC, ex2_intro/ +/5 width=11 by sex_inv_tc_dx, sex_co, ext2_inv_tc, ext2_refl, monotonic_TC, ex2_intro/ qed-. lemma s_rs_transitive_lex_inv_isid (R): s_rs_transitive … R (λ_.lex R) → @@ -46,9 +46,9 @@ qed-. lemma lex_CTC (R): s_rs_transitive … R (λ_. lex R) → TC … (lex R) ⊆ lex (CTC … R). #R #HR #L1 #L2 #HL12 -lapply (monotonic_TC … (lexs cfull (cext2 R) 𝐈𝐝) … HL12) -HL12 -[ #L1 #L2 * /3 width=3 by lexs_eq_repl_fwd, eq_id_inv_isid/ -| /5 width=9 by s_rs_transitive_lex_inv_isid, lexs_tc_dx, lexs_co, ext2_tc, ex2_intro/ +lapply (monotonic_TC … (sex cfull (cext2 R) 𝐈𝐝) … HL12) -HL12 +[ #L1 #L2 * /3 width=3 by sex_eq_repl_fwd, eq_id_inv_isid/ +| /5 width=9 by s_rs_transitive_lex_inv_isid, sex_tc_dx, sex_co, ext2_tc, ex2_intro/ ] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lexs.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lexs.ma deleted file mode 100644 index 0ff659b13..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lexs.ma +++ /dev/null @@ -1,296 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "ground_2/relocation/rtmap_sle.ma". -include "ground_2/relocation/rtmap_sdj.ma". -include "basic_2/notation/relations/relationstar_5.ma". -include "basic_2/syntax/lenv.ma". - -(* GENERIC ENTRYWISE EXTENSION OF CONTEXT-SENSITIVE REALTIONS FOR TERMS *****) - -inductive lexs (RN,RP:relation3 lenv bind bind): rtmap → relation lenv ≝ -| lexs_atom: ∀f. lexs RN RP f (⋆) (⋆) -| lexs_next: ∀f,I1,I2,L1,L2. - lexs RN RP f L1 L2 → RN L1 I1 I2 → - lexs RN RP (↑f) (L1.ⓘ{I1}) (L2.ⓘ{I2}) -| lexs_push: ∀f,I1,I2,L1,L2. - lexs RN RP f L1 L2 → RP L1 I1 I2 → - lexs RN RP (⫯f) (L1.ⓘ{I1}) (L2.ⓘ{I2}) -. - -interpretation "generic entrywise extension (local environment)" - 'RelationStar RN RP f L1 L2 = (lexs RN RP f L1 L2). - -definition R_pw_confluent2_lexs: relation3 lenv bind bind → relation3 lenv bind bind → - relation3 lenv bind bind → relation3 lenv bind bind → - relation3 lenv bind bind → relation3 lenv bind bind → - relation3 rtmap lenv bind ≝ - λR1,R2,RN1,RP1,RN2,RP2,f,L0,I0. - ∀I1. R1 L0 I0 I1 → ∀I2. R2 L0 I0 I2 → - ∀L1. L0 ⪤*[RN1, RP1, f] L1 → ∀L2. L0 ⪤*[RN2, RP2, f] L2 → - ∃∃I. R2 L1 I1 I & R1 L2 I2 I. - -definition lexs_transitive: relation3 lenv bind bind → relation3 lenv bind bind → - relation3 lenv bind bind → - relation3 lenv bind bind → relation3 lenv bind bind → - relation3 rtmap lenv bind ≝ - λR1,R2,R3,RN,RP,f,L1,I1. - ∀I. R1 L1 I1 I → ∀L2. L1 ⪤*[RN, RP, f] L2 → - ∀I2. R2 L2 I I2 → R3 L1 I1 I2. - -(* Basic inversion lemmas ***************************************************) - -fact lexs_inv_atom1_aux: ∀RN,RP,f,X,Y. X ⪤*[RN, RP, f] Y → X = ⋆ → Y = ⋆. -#RN #RP #f #X #Y * -f -X -Y // -#f #I1 #I2 #L1 #L2 #_ #_ #H destruct -qed-. - -(* Basic_2A1: includes lpx_sn_inv_atom1 *) -lemma lexs_inv_atom1: ∀RN,RP,f,Y. ⋆ ⪤*[RN, RP, f] Y → Y = ⋆. -/2 width=6 by lexs_inv_atom1_aux/ qed-. - -fact lexs_inv_next1_aux: ∀RN,RP,f,X,Y. X ⪤*[RN, RP, f] Y → ∀g,J1,K1. X = K1.ⓘ{J1} → f = ↑g → - ∃∃J2,K2. K1 ⪤*[RN, RP, g] K2 & RN K1 J1 J2 & Y = K2.ⓘ{J2}. -#RN #RP #f #X #Y * -f -X -Y -[ #f #g #J1 #K1 #H destruct -| #f #I1 #I2 #L1 #L2 #HL #HI #g #J1 #K1 #H1 #H2 <(injective_next … H2) -g destruct - /2 width=5 by ex3_2_intro/ -| #f #I1 #I2 #L1 #L2 #_ #_ #g #J1 #K1 #_ #H elim (discr_push_next … H) -] -qed-. - -(* Basic_2A1: includes lpx_sn_inv_pair1 *) -lemma lexs_inv_next1: ∀RN,RP,g,J1,K1,Y. K1.ⓘ{J1} ⪤*[RN, RP, ↑g] Y → - ∃∃J2,K2. K1 ⪤*[RN, RP, g] K2 & RN K1 J1 J2 & Y = K2.ⓘ{J2}. -/2 width=7 by lexs_inv_next1_aux/ qed-. - -fact lexs_inv_push1_aux: ∀RN,RP,f,X,Y. X ⪤*[RN, RP, f] Y → ∀g,J1,K1. X = K1.ⓘ{J1} → f = ⫯g → - ∃∃J2,K2. K1 ⪤*[RN, RP, g] K2 & RP K1 J1 J2 & Y = K2.ⓘ{J2}. -#RN #RP #f #X #Y * -f -X -Y -[ #f #g #J1 #K1 #H destruct -| #f #I1 #I2 #L1 #L2 #_ #_ #g #J1 #K1 #_ #H elim (discr_next_push … H) -| #f #I1 #I2 #L1 #L2 #HL #HI #g #J1 #K1 #H1 #H2 <(injective_push … H2) -g destruct - /2 width=5 by ex3_2_intro/ -] -qed-. - -lemma lexs_inv_push1: ∀RN,RP,g,J1,K1,Y. K1.ⓘ{J1} ⪤*[RN, RP, ⫯g] Y → - ∃∃J2,K2. K1 ⪤*[RN, RP, g] K2 & RP K1 J1 J2 & Y = K2.ⓘ{J2}. -/2 width=7 by lexs_inv_push1_aux/ qed-. - -fact lexs_inv_atom2_aux: ∀RN,RP,f,X,Y. X ⪤*[RN, RP, f] Y → Y = ⋆ → X = ⋆. -#RN #RP #f #X #Y * -f -X -Y // -#f #I1 #I2 #L1 #L2 #_ #_ #H destruct -qed-. - -(* Basic_2A1: includes lpx_sn_inv_atom2 *) -lemma lexs_inv_atom2: ∀RN,RP,f,X. X ⪤*[RN, RP, f] ⋆ → X = ⋆. -/2 width=6 by lexs_inv_atom2_aux/ qed-. - -fact lexs_inv_next2_aux: ∀RN,RP,f,X,Y. X ⪤*[RN, RP, f] Y → ∀g,J2,K2. Y = K2.ⓘ{J2} → f = ↑g → - ∃∃J1,K1. K1 ⪤*[RN, RP, g] K2 & RN K1 J1 J2 & X = K1.ⓘ{J1}. -#RN #RP #f #X #Y * -f -X -Y -[ #f #g #J2 #K2 #H destruct -| #f #I1 #I2 #L1 #L2 #HL #HI #g #J2 #K2 #H1 #H2 <(injective_next … H2) -g destruct - /2 width=5 by ex3_2_intro/ -| #f #I1 #I2 #L1 #L2 #_ #_ #g #J2 #K2 #_ #H elim (discr_push_next … H) -] -qed-. - -(* Basic_2A1: includes lpx_sn_inv_pair2 *) -lemma lexs_inv_next2: ∀RN,RP,g,J2,X,K2. X ⪤*[RN, RP, ↑g] K2.ⓘ{J2} → - ∃∃J1,K1. K1 ⪤*[RN, RP, g] K2 & RN K1 J1 J2 & X = K1.ⓘ{J1}. -/2 width=7 by lexs_inv_next2_aux/ qed-. - -fact lexs_inv_push2_aux: ∀RN,RP,f,X,Y. X ⪤*[RN, RP, f] Y → ∀g,J2,K2. Y = K2.ⓘ{J2} → f = ⫯g → - ∃∃J1,K1. K1 ⪤*[RN, RP, g] K2 & RP K1 J1 J2 & X = K1.ⓘ{J1}. -#RN #RP #f #X #Y * -f -X -Y -[ #f #J2 #K2 #g #H destruct -| #f #I1 #I2 #L1 #L2 #_ #_ #g #J2 #K2 #_ #H elim (discr_next_push … H) -| #f #I1 #I2 #L1 #L2 #HL #HI #g #J2 #K2 #H1 #H2 <(injective_push … H2) -g destruct - /2 width=5 by ex3_2_intro/ -] -qed-. - -lemma lexs_inv_push2: ∀RN,RP,g,J2,X,K2. X ⪤*[RN, RP, ⫯g] K2.ⓘ{J2} → - ∃∃J1,K1. K1 ⪤*[RN, RP, g] K2 & RP K1 J1 J2 & X = K1.ⓘ{J1}. -/2 width=7 by lexs_inv_push2_aux/ qed-. - -(* Basic_2A1: includes lpx_sn_inv_pair *) -lemma lexs_inv_next: ∀RN,RP,f,I1,I2,L1,L2. - L1.ⓘ{I1} ⪤*[RN, RP, ↑f] L2.ⓘ{I2} → - L1 ⪤*[RN, RP, f] L2 ∧ RN L1 I1 I2. -#RN #RP #f #I1 #I2 #L1 #L2 #H elim (lexs_inv_next1 … H) -H -#I0 #L0 #HL10 #HI10 #H destruct /2 width=1 by conj/ -qed-. - -lemma lexs_inv_push: ∀RN,RP,f,I1,I2,L1,L2. - L1.ⓘ{I1} ⪤*[RN, RP, ⫯f] L2.ⓘ{I2} → - L1 ⪤*[RN, RP, f] L2 ∧ RP L1 I1 I2. -#RN #RP #f #I1 #I2 #L1 #L2 #H elim (lexs_inv_push1 … H) -H -#I0 #L0 #HL10 #HI10 #H destruct /2 width=1 by conj/ -qed-. - -lemma lexs_inv_tl: ∀RN,RP,f,I1,I2,L1,L2. L1 ⪤*[RN, RP, ⫱f] L2 → - RN L1 I1 I2 → RP L1 I1 I2 → - L1.ⓘ{I1} ⪤*[RN, RP, f] L2.ⓘ{I2}. -#RN #RP #f #I1 #I2 #L2 #L2 elim (pn_split f) * -/2 width=1 by lexs_next, lexs_push/ -qed-. - -(* Basic forward lemmas *****************************************************) - -lemma lexs_fwd_bind: ∀RN,RP,f,I1,I2,L1,L2. - L1.ⓘ{I1} ⪤*[RN, RP, f] L2.ⓘ{I2} → - L1 ⪤*[RN, RP, ⫱f] L2. -#RN #RP #f #I1 #I2 #L1 #L2 #Hf -elim (pn_split f) * #g #H destruct -[ elim (lexs_inv_push … Hf) | elim (lexs_inv_next … Hf) ] -Hf // -qed-. - -(* Basic properties *********************************************************) - -lemma lexs_eq_repl_back: ∀RN,RP,L1,L2. eq_repl_back … (λf. L1 ⪤*[RN, RP, f] L2). -#RN #RP #L1 #L2 #f1 #H elim H -f1 -L1 -L2 // -#f1 #I1 #I2 #L1 #L2 #_ #HI #IH #f2 #H -[ elim (eq_inv_nx … H) -H /3 width=3 by lexs_next/ -| elim (eq_inv_px … H) -H /3 width=3 by lexs_push/ -] -qed-. - -lemma lexs_eq_repl_fwd: ∀RN,RP,L1,L2. eq_repl_fwd … (λf. L1 ⪤*[RN, RP, f] L2). -#RN #RP #L1 #L2 @eq_repl_sym /2 width=3 by lexs_eq_repl_back/ (**) (* full auto fails *) -qed-. - -lemma lexs_refl: ∀RN,RP. c_reflexive … RN → c_reflexive … RP → - ∀f.reflexive … (lexs RN RP f). -#RN #RP #HRN #HRP #f #L generalize in match f; -f elim L -L // -#L #I #IH #f elim (pn_split f) * -#g #H destruct /2 width=1 by lexs_next, lexs_push/ -qed. - -lemma lexs_sym: ∀RN,RP. - (∀L1,L2,I1,I2. RN L1 I1 I2 → RN L2 I2 I1) → - (∀L1,L2,I1,I2. RP L1 I1 I2 → RP L2 I2 I1) → - ∀f. symmetric … (lexs RN RP f). -#RN #RP #HRN #HRP #f #L1 #L2 #H elim H -L1 -L2 -f -/3 width=2 by lexs_next, lexs_push/ -qed-. - -lemma lexs_pair_repl: ∀RN,RP,f,I1,I2,L1,L2. - L1.ⓘ{I1} ⪤*[RN, RP, f] L2.ⓘ{I2} → - ∀J1,J2. RN L1 J1 J2 → RP L1 J1 J2 → - L1.ⓘ{J1} ⪤*[RN, RP, f] L2.ⓘ{J2}. -/3 width=3 by lexs_inv_tl, lexs_fwd_bind/ qed-. - -lemma lexs_co: ∀RN1,RP1,RN2,RP2. RN1 ⊆ RN2 → RP1 ⊆ RP2 → - ∀f,L1,L2. L1 ⪤*[RN1, RP1, f] L2 → L1 ⪤*[RN2, RP2, f] L2. -#RN1 #RP1 #RN2 #RP2 #HRN #HRP #f #L1 #L2 #H elim H -f -L1 -L2 -/3 width=1 by lexs_atom, lexs_next, lexs_push/ -qed-. - -lemma lexs_co_isid: ∀RN1,RP1,RN2,RP2. RP1 ⊆ RP2 → - ∀f,L1,L2. L1 ⪤*[RN1, RP1, f] L2 → 𝐈⦃f⦄ → - L1 ⪤*[RN2, RP2, f] L2. -#RN1 #RP1 #RN2 #RP2 #HR #f #L1 #L2 #H elim H -f -L1 -L2 // -#f #I1 #I2 #K1 #K2 #_ #HI12 #IH #H -[ elim (isid_inv_next … H) -H // -| /4 width=3 by lexs_push, isid_inv_push/ -] -qed-. - -lemma lexs_sdj: ∀RN,RP. RP ⊆ RN → - ∀f1,L1,L2. L1 ⪤*[RN, RP, f1] L2 → - ∀f2. f1 ∥ f2 → L1 ⪤*[RP, RN, f2] L2. -#RN #RP #HR #f1 #L1 #L2 #H elim H -f1 -L1 -L2 // -#f1 #I1 #I2 #L1 #L2 #_ #HI12 #IH #f2 #H12 -[ elim (sdj_inv_nx … H12) -H12 [2,3: // ] - #g2 #H #H2 destruct /3 width=1 by lexs_push/ -| elim (sdj_inv_px … H12) -H12 [2,4: // ] * - #g2 #H #H2 destruct /3 width=1 by lexs_next, lexs_push/ -] -qed-. - -lemma sle_lexs_trans: ∀RN,RP. RN ⊆ RP → - ∀f2,L1,L2. L1 ⪤*[RN, RP, f2] L2 → - ∀f1. f1 ⊆ f2 → L1 ⪤*[RN, RP, f1] L2. -#RN #RP #HR #f2 #L1 #L2 #H elim H -f2 -L1 -L2 // -#f2 #I1 #I2 #L1 #L2 #_ #HI12 #IH #f1 #H12 -[ elim (pn_split f1) * ] -[ /4 width=5 by lexs_push, sle_inv_pn/ -| /4 width=5 by lexs_next, sle_inv_nn/ -| elim (sle_inv_xp … H12) -H12 [2,3: // ] - #g1 #H #H1 destruct /3 width=5 by lexs_push/ -] -qed-. - -lemma sle_lexs_conf: ∀RN,RP. RP ⊆ RN → - ∀f1,L1,L2. L1 ⪤*[RN, RP, f1] L2 → - ∀f2. f1 ⊆ f2 → L1 ⪤*[RN, RP, f2] L2. -#RN #RP #HR #f1 #L1 #L2 #H elim H -f1 -L1 -L2 // -#f1 #I1 #I2 #L1 #L2 #_ #HI12 #IH #f2 #H12 -[2: elim (pn_split f2) * ] -[ /4 width=5 by lexs_push, sle_inv_pp/ -| /4 width=5 by lexs_next, sle_inv_pn/ -| elim (sle_inv_nx … H12) -H12 [2,3: // ] - #g2 #H #H2 destruct /3 width=5 by lexs_next/ -] -qed-. - -lemma lexs_sle_split: ∀R1,R2,RP. c_reflexive … R1 → c_reflexive … R2 → - ∀f,L1,L2. L1 ⪤*[R1, RP, f] L2 → ∀g. f ⊆ g → - ∃∃L. L1 ⪤*[R1, RP, g] L & L ⪤*[R2, cfull, f] L2. -#R1 #R2 #RP #HR1 #HR2 #f #L1 #L2 #H elim H -f -L1 -L2 -[ /2 width=3 by lexs_atom, ex2_intro/ ] -#f #I1 #I2 #L1 #L2 #_ #HI12 #IH #y #H -[ elim (sle_inv_nx … H ??) -H [ |*: // ] #g #Hfg #H destruct - elim (IH … Hfg) -IH -Hfg /3 width=5 by lexs_next, ex2_intro/ -| elim (sle_inv_px … H ??) -H [1,3: * |*: // ] #g #Hfg #H destruct - elim (IH … Hfg) -IH -Hfg /3 width=5 by lexs_next, lexs_push, ex2_intro/ -] -qed-. - -lemma lexs_sdj_split: ∀R1,R2,RP. c_reflexive … R1 → c_reflexive … R2 → - ∀f,L1,L2. L1 ⪤*[R1, RP, f] L2 → ∀g. f ∥ g → - ∃∃L. L1 ⪤*[RP, R1, g] L & L ⪤*[R2, cfull, f] L2. -#R1 #R2 #RP #HR1 #HR2 #f #L1 #L2 #H elim H -f -L1 -L2 -[ /2 width=3 by lexs_atom, ex2_intro/ ] -#f #I1 #I2 #L1 #L2 #_ #HI12 #IH #y #H -[ elim (sdj_inv_nx … H ??) -H [ |*: // ] #g #Hfg #H destruct - elim (IH … Hfg) -IH -Hfg /3 width=5 by lexs_next, lexs_push, ex2_intro/ -| elim (sdj_inv_px … H ??) -H [1,3: * |*: // ] #g #Hfg #H destruct - elim (IH … Hfg) -IH -Hfg /3 width=5 by lexs_next, lexs_push, ex2_intro/ -] -qed-. - -lemma lexs_dec: ∀RN,RP. - (∀L,I1,I2. Decidable (RN L I1 I2)) → - (∀L,I1,I2. Decidable (RP L I1 I2)) → - ∀L1,L2,f. Decidable (L1 ⪤*[RN, RP, f] L2). -#RN #RP #HRN #HRP #L1 elim L1 -L1 [ * | #L1 #I1 #IH * ] -[ /2 width=1 by lexs_atom, or_introl/ -| #L2 #I2 #f @or_intror #H - lapply (lexs_inv_atom1 … H) -H #H destruct -| #f @or_intror #H - lapply (lexs_inv_atom2 … H) -H #H destruct -| #L2 #I2 #f elim (IH L2 (⫱f)) -IH #HL12 - [2: /4 width=3 by lexs_fwd_bind, or_intror/ ] - elim (pn_split f) * #g #H destruct - [ elim (HRP L1 I1 I2) | elim (HRN L1 I1 I2) ] -HRP -HRN #HV12 - [1,3: /3 width=1 by lexs_push, lexs_next, or_introl/ ] - @or_intror #H - [ elim (lexs_inv_push … H) | elim (lexs_inv_next … H) ] -H - /2 width=1 by/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lexs_length.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lexs_length.ma deleted file mode 100644 index 71467ba72..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lexs_length.ma +++ /dev/null @@ -1,46 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/syntax/lenv_length.ma". -include "basic_2/relocation/lexs.ma". - -(* GENERIC ENTRYWISE EXTENSION OF CONTEXT-SENSITIVE REALTIONS FOR TERMS *****) - -(* Forward lemmas with length for local environments ************************) - -lemma lexs_fwd_length: ∀RN,RP,f,L1,L2. L1 ⪤*[RN, RP, f] L2 → |L1| = |L2|. -#RN #RP #f #L1 #L2 #H elim H -f -L1 -L2 // -#f #I1 #I2 #L1 #L2 >length_bind >length_bind // -qed-. - -(* Properties with length for local environments ****************************) - -lemma lexs_length_cfull: ∀L1,L2. |L1| = |L2| → ∀f. L1 ⪤*[cfull, cfull, f] L2. -#L1 elim L1 -L1 -[ #Y2 #H >(length_inv_zero_sn … H) -Y2 // -| #L1 #I1 #IH #Y2 #H #f - elim (length_inv_succ_sn … H) -H #I2 #L2 #HL12 #H destruct - elim (pn_split f) * #g #H destruct /3 width=1 by lexs_next, lexs_push/ -] -qed. - -lemma lexs_length_isid: ∀R,L1,L2. |L1| = |L2| → - ∀f. 𝐈⦃f⦄ → L1 ⪤*[R, cfull, f] L2. -#R #L1 elim L1 -L1 -[ #Y2 #H >(length_inv_zero_sn … H) -Y2 // -| #L1 #I1 #IH #Y2 #H #f #Hf - elim (length_inv_succ_sn … H) -H #I2 #L2 #HL12 #H destruct - elim (isid_inv_gen … Hf) -Hf #g #Hg #H destruct /3 width=1 by lexs_push/ -] -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lexs_lexs.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lexs_lexs.ma deleted file mode 100644 index 753b43b2e..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lexs_lexs.ma +++ /dev/null @@ -1,119 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "ground_2/relocation/rtmap_sand.ma". -include "basic_2/relocation/drops.ma". - -(* GENERIC ENTRYWISE EXTENSION OF CONTEXT-SENSITIVE REALTIONS FOR TERMS *****) - -(* Main properties **********************************************************) - -theorem lexs_trans_gen (RN1) (RP1) (RN2) (RP2) (RN) (RP): - ∀L1,f. - (∀g,I,K,n. ⬇*[n] L1 ≘ K.ⓘ{I} → ↑g = ⫱*[n] f → lexs_transitive RN1 RN2 RN RN1 RP1 g K I) → - (∀g,I,K,n. ⬇*[n] L1 ≘ K.ⓘ{I} → ⫯g = ⫱*[n] f → lexs_transitive RP1 RP2 RP RN1 RP1 g K I) → - ∀L0. L1 ⪤*[RN1, RP1, f] L0 → - ∀L2. L0 ⪤*[RN2, RP2, f] L2 → - L1 ⪤*[RN, RP, f] L2. -#RN1 #RP1 #RN2 #RP2 #RN #RP #L1 elim L1 -L1 -[ #f #_ #_ #L0 #H1 #L2 #H2 - lapply (lexs_inv_atom1 … H1) -H1 #H destruct - lapply (lexs_inv_atom1 … H2) -H2 #H destruct - /2 width=1 by lexs_atom/ -| #K1 #I1 #IH #f elim (pn_split f) * #g #H destruct - #HN #HP #L0 #H1 #L2 #H2 - [ elim (lexs_inv_push1 … H1) -H1 #I0 #K0 #HK10 #HI10 #H destruct - elim (lexs_inv_push1 … H2) -H2 #I2 #K2 #HK02 #HI02 #H destruct - lapply (HP … 0 … HI10 … HK10 … HI02) -HI10 -HI02 /2 width=2 by drops_refl/ #HI12 - lapply (IH … HK10 … HK02) -IH -K0 /3 width=3 by lexs_push, drops_drop/ - | elim (lexs_inv_next1 … H1) -H1 #I0 #K0 #HK10 #HI10 #H destruct - elim (lexs_inv_next1 … H2) -H2 #I2 #K2 #HK02 #HI02 #H destruct - lapply (HN … 0 … HI10 … HK10 … HI02) -HI10 -HI02 /2 width=2 by drops_refl/ #HI12 - lapply (IH … HK10 … HK02) -IH -K0 /3 width=3 by lexs_next, drops_drop/ - ] -] -qed-. - -theorem lexs_trans (RN) (RP) (f): (∀g,I,K. lexs_transitive RN RN RN RN RP g K I) → - (∀g,I,K. lexs_transitive RP RP RP RN RP g K I) → - Transitive … (lexs RN RP f). -/2 width=9 by lexs_trans_gen/ qed-. - -theorem lexs_trans_id_cfull: ∀R1,R2,R3,L1,L,f. L1 ⪤*[R1, cfull, f] L → 𝐈⦃f⦄ → - ∀L2. L ⪤*[R2, cfull, f] L2 → L1 ⪤*[R3, cfull, f] L2. -#R1 #R2 #R3 #L1 #L #f #H elim H -L1 -L -f -[ #f #Hf #L2 #H >(lexs_inv_atom1 … H) -L2 // ] -#f #I1 #I #K1 #K #HK1 #_ #IH #Hf #L2 #H -[ elim (isid_inv_next … Hf) | lapply (isid_inv_push … Hf ??) ] -Hf [5: |*: // ] #Hf -elim (lexs_inv_push1 … H) -H #I2 #K2 #HK2 #_ #H destruct -/3 width=1 by lexs_push/ -qed-. - -theorem lexs_conf (RN1) (RP1) (RN2) (RP2): - ∀L,f. - (∀g,I,K,n. ⬇*[n] L ≘ K.ⓘ{I} → ↑g = ⫱*[n] f → R_pw_confluent2_lexs RN1 RN2 RN1 RP1 RN2 RP2 g K I) → - (∀g,I,K,n. ⬇*[n] L ≘ K.ⓘ{I} → ⫯g = ⫱*[n] f → R_pw_confluent2_lexs RP1 RP2 RN1 RP1 RN2 RP2 g K I) → - pw_confluent2 … (lexs RN1 RP1 f) (lexs RN2 RP2 f) L. -#RN1 #RP1 #RN2 #RP2 #L elim L -L -[ #f #_ #_ #L1 #H1 #L2 #H2 >(lexs_inv_atom1 … H1) >(lexs_inv_atom1 … H2) -H2 -H1 - /2 width=3 by lexs_atom, ex2_intro/ -| #L #I0 #IH #f elim (pn_split f) * #g #H destruct - #HN #HP #Y1 #H1 #Y2 #H2 - [ elim (lexs_inv_push1 … H1) -H1 #I1 #L1 #HL1 #HI01 #H destruct - elim (lexs_inv_push1 … H2) -H2 #I2 #L2 #HL2 #HI02 #H destruct - elim (HP … 0 … HI01 … HI02 … HL1 … HL2) -HI01 -HI02 /2 width=2 by drops_refl/ #I #HI1 #HI2 - elim (IH … HL1 … HL2) -IH -HL1 -HL2 /3 width=5 by drops_drop, lexs_push, ex2_intro/ - | elim (lexs_inv_next1 … H1) -H1 #I1 #L1 #HL1 #HI01 #H destruct - elim (lexs_inv_next1 … H2) -H2 #I2 #L2 #HL2 #HI02 #H destruct - elim (HN … 0 … HI01 … HI02 … HL1 … HL2) -HI01 -HI02 /2 width=2 by drops_refl/ #I #HI1 #HI2 - elim (IH … HL1 … HL2) -IH -HL1 -HL2 /3 width=5 by drops_drop, lexs_next, ex2_intro/ - ] -] -qed-. - -theorem lexs_canc_sn: ∀RN,RP,f. Transitive … (lexs RN RP f) → - symmetric … (lexs RN RP f) → - left_cancellable … (lexs RN RP f). -/3 width=3 by/ qed-. - -theorem lexs_canc_dx: ∀RN,RP,f. Transitive … (lexs RN RP f) → - symmetric … (lexs RN RP f) → - right_cancellable … (lexs RN RP f). -/3 width=3 by/ qed-. - -lemma lexs_meet: ∀RN,RP,L1,L2. - ∀f1. L1 ⪤*[RN, RP, f1] L2 → - ∀f2. L1 ⪤*[RN, RP, f2] L2 → - ∀f. f1 ⋒ f2 ≘ f → L1 ⪤*[RN, RP, f] L2. -#RN #RP #L1 #L2 #f1 #H elim H -f1 -L1 -L2 // -#f1 #I1 #I2 #L1 #L2 #_ #HI12 #IH #f2 #H #f #Hf -elim (pn_split f2) * #g2 #H2 destruct -try elim (lexs_inv_push … H) try elim (lexs_inv_next … H) -H -[ elim (sand_inv_npx … Hf) | elim (sand_inv_nnx … Hf) -| elim (sand_inv_ppx … Hf) | elim (sand_inv_pnx … Hf) -] -Hf /3 width=5 by lexs_next, lexs_push/ -qed-. - -lemma lexs_join: ∀RN,RP,L1,L2. - ∀f1. L1 ⪤*[RN, RP, f1] L2 → - ∀f2. L1 ⪤*[RN, RP, f2] L2 → - ∀f. f1 ⋓ f2 ≘ f → L1 ⪤*[RN, RP, f] L2. -#RN #RP #L1 #L2 #f1 #H elim H -f1 -L1 -L2 // -#f1 #I1 #I2 #L1 #L2 #_ #HI12 #IH #f2 #H #f #Hf -elim (pn_split f2) * #g2 #H2 destruct -try elim (lexs_inv_push … H) try elim (lexs_inv_next … H) -H -[ elim (sor_inv_npx … Hf) | elim (sor_inv_nnx … Hf) -| elim (sor_inv_ppx … Hf) | elim (sor_inv_pnx … Hf) -] -Hf /3 width=5 by lexs_next, lexs_push/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lexs_tc.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lexs_tc.ma deleted file mode 100644 index 0ddcf6eb2..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lexs_tc.ma +++ /dev/null @@ -1,119 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "ground_2/lib/star.ma". -include "basic_2/relocation/lexs.ma". - -(* GENERIC ENTRYWISE EXTENSION OF CONTEXT-SENSITIVE REALTIONS FOR TERMS *****) - -definition s_rs_transitive_isid: relation (relation3 lenv bind bind) ≝ λRN,RP. - ∀f. 𝐈⦃f⦄ → s_rs_transitive … RP (λ_.lexs RN RP f). - -(* Properties with transitive closure ***************************************) - -lemma lexs_tc_refl: ∀RN,RP. c_reflexive … RN → c_reflexive … RP → - ∀f. reflexive … (TC … (lexs RN RP f)). -/3 width=1 by lexs_refl, TC_reflexive/ qed. - -lemma lexs_tc_next_sn: ∀RN,RP. c_reflexive … RN → - ∀f,I2,L1,L2. TC … (lexs RN RP f) L1 L2 → ∀I1. RN L1 I1 I2 → - TC … (lexs RN RP (↑f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}). -#RN #RP #HRN #f #I2 #L1 #L2 #H @(TC_ind_dx ??????? H) -L1 -/3 width=3 by lexs_next, TC_strap, inj/ -qed. - -lemma lexs_tc_next_dx: ∀RN,RP. c_reflexive … RN → c_reflexive … RP → - ∀f,I1,I2,L1. (CTC … RN) L1 I1 I2 → ∀L2. L1 ⪤*[RN, RP, f] L2 → - TC … (lexs RN RP (↑f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}). -#RN #RP #HRN #HRP #f #I1 #I2 #L1 #H elim H -I2 -/4 width=5 by lexs_refl, lexs_next, step, inj/ -qed. - -lemma lexs_tc_push_sn: ∀RN,RP. c_reflexive … RP → - ∀f,I2,L1,L2. TC … (lexs RN RP f) L1 L2 → ∀I1. RP L1 I1 I2 → - TC … (lexs RN RP (⫯f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}). -#RN #RP #HRP #f #I2 #L1 #L2 #H @(TC_ind_dx ??????? H) -L1 -/3 width=3 by lexs_push, TC_strap, inj/ -qed. - -lemma lexs_tc_push_dx: ∀RN,RP. c_reflexive … RN → c_reflexive … RP → - ∀f,I1,I2,L1. (CTC … RP) L1 I1 I2 → ∀L2. L1 ⪤*[RN, RP, f] L2 → - TC … (lexs RN RP (⫯f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}). -#RN #RP #HRN #HRP #f #I1 #I2 #L1 #H elim H -I2 -/4 width=5 by lexs_refl, lexs_push, step, inj/ -qed. - -lemma lexs_tc_inj_sn: ∀RN,RP,f,L1,L2. L1 ⪤*[RN, RP, f] L2 → L1 ⪤*[CTC … RN, RP, f] L2. -#RN #RP #f #L1 #L2 #H elim H -f -L1 -L2 -/3 width=1 by lexs_push, lexs_next, inj/ -qed. - -lemma lexs_tc_inj_dx: ∀RN,RP,f,L1,L2. L1 ⪤*[RN, RP, f] L2 → L1 ⪤*[RN, CTC … RP, f] L2. -#RN #RP #f #L1 #L2 #H elim H -f -L1 -L2 -/3 width=1 by lexs_push, lexs_next, inj/ -qed. - -(* Main properties with transitive closure **********************************) - -theorem lexs_tc_next: ∀RN,RP. c_reflexive … RN → c_reflexive … RP → - ∀f,I1,I2,L1. (CTC … RN) L1 I1 I2 → ∀L2. TC … (lexs RN RP f) L1 L2 → - TC … (lexs RN RP (↑f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}). -#RN #RP #HRN #HRP #f #I1 #I2 #L1 #H elim H -I2 -/4 width=5 by lexs_tc_next_sn, lexs_tc_refl, trans_TC/ -qed. - -theorem lexs_tc_push: ∀RN,RP. c_reflexive … RN → c_reflexive … RP → - ∀f,I1,I2,L1. (CTC … RP) L1 I1 I2 → ∀L2. TC … (lexs RN RP f) L1 L2 → - TC … (lexs RN RP (⫯f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}). -#RN #RP #HRN #HRP #f #I1 #I2 #L1 #H elim H -I2 -/4 width=5 by lexs_tc_push_sn, lexs_tc_refl, trans_TC/ -qed. - -(* Basic_2A1: uses: TC_lpx_sn_ind *) -theorem lexs_tc_step_dx: ∀RN,RP. s_rs_transitive_isid RN RP → - ∀f,L1,L. L1 ⪤*[RN, RP, f] L → 𝐈⦃f⦄ → - ∀L2. L ⪤*[RN, CTC … RP, f] L2 → L1⪤* [RN, CTC … RP, f] L2. -#RN #RP #HRP #f #L1 #L #H elim H -f -L1 -L -[ #f #_ #Y #H -HRP >(lexs_inv_atom1 … H) -Y // ] -#f #I1 #I #L1 #L #HL1 #HI1 #IH #Hf #Y #H -[ elim (isid_inv_next … Hf) -Hf // -| lapply (isid_inv_push … Hf ??) -Hf [3: |*: // ] #Hf - elim (lexs_inv_push1 … H) -H #I2 #L2 #HL2 #HI2 #H destruct - @lexs_push [ /2 width=1 by/ ] -L2 -IH - @(TC_strap … HI1) -HI1 - @(HRP … HL1) // (**) (* auto fails *) -] -qed-. - -(* Advanced properties ******************************************************) - -lemma lexs_tc_dx: ∀RN,RP. s_rs_transitive_isid RN RP → - ∀f. 𝐈⦃f⦄ → ∀L1,L2. TC … (lexs RN RP f) L1 L2 → L1 ⪤*[RN, CTC … RP, f] L2. -#RN #RP #HRP #f #Hf #L1 #L2 #H @(TC_ind_dx ??????? H) -L1 -/3 width=3 by lexs_tc_step_dx, lexs_tc_inj_dx/ -qed. - -(* Advanced inversion lemmas ************************************************) - -lemma lexs_inv_tc_sn: ∀RN,RP. c_reflexive … RN → c_reflexive … RP → - ∀f,L1,L2. L1 ⪤*[CTC … RN, RP, f] L2 → TC … (lexs RN RP f) L1 L2. -#RN #RP #HRN #HRP #f #L1 #L2 #H elim H -f -L1 -L2 -/2 width=1 by lexs_tc_next, lexs_tc_push_sn, lexs_atom, inj/ -qed-. - -lemma lexs_inv_tc_dx: ∀RN,RP. c_reflexive … RN → c_reflexive … RP → - ∀f,L1,L2. L1 ⪤*[RN, CTC … RP, f] L2 → TC … (lexs RN RP f) L1 L2. -#RN #RP #HRN #HRP #f #L1 #L2 #H elim H -f -L1 -L2 -/2 width=1 by lexs_tc_push, lexs_tc_next_sn, lexs_atom, inj/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lreq.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lreq.ma deleted file mode 100644 index 747e9a511..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lreq.ma +++ /dev/null @@ -1,103 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/notation/relations/ideqsn_3.ma". -include "basic_2/syntax/ceq_ext.ma". -include "basic_2/relocation/lexs.ma". - -(* RANGED EQUIVALENCE FOR LOCAL ENVIRONMENTS ********************************) - -(* Basic_2A1: includes: lreq_atom lreq_zero lreq_pair lreq_succ *) -definition lreq: relation3 rtmap lenv lenv ≝ lexs ceq_ext cfull. - -interpretation - "ranged equivalence (local environment)" - 'IdEqSn f L1 L2 = (lreq f L1 L2). - -(* Basic properties *********************************************************) - -lemma lreq_eq_repl_back: ∀L1,L2. eq_repl_back … (λf. L1 ≡[f] L2). -/2 width=3 by lexs_eq_repl_back/ qed-. - -lemma lreq_eq_repl_fwd: ∀L1,L2. eq_repl_fwd … (λf. L1 ≡[f] L2). -/2 width=3 by lexs_eq_repl_fwd/ qed-. - -lemma sle_lreq_trans: ∀f2,L1,L2. L1 ≡[f2] L2 → - ∀f1. f1 ⊆ f2 → L1 ≡[f1] L2. -/2 width=3 by sle_lexs_trans/ qed-. - -(* Basic_2A1: includes: lreq_refl *) -lemma lreq_refl: ∀f. reflexive … (lreq f). -/2 width=1 by lexs_refl/ qed. - -(* Basic_2A1: includes: lreq_sym *) -lemma lreq_sym: ∀f. symmetric … (lreq f). -/3 width=2 by lexs_sym, cext2_sym/ qed-. - -(* Basic inversion lemmas ***************************************************) - -(* Basic_2A1: includes: lreq_inv_atom1 *) -lemma lreq_inv_atom1: ∀f,Y. ⋆ ≡[f] Y → Y = ⋆. -/2 width=4 by lexs_inv_atom1/ qed-. - -(* Basic_2A1: includes: lreq_inv_pair1 *) -lemma lreq_inv_next1: ∀g,J,K1,Y. K1.ⓘ{J} ≡[↑g] Y → - ∃∃K2. K1 ≡[g] K2 & Y = K2.ⓘ{J}. -#g #J #K1 #Y #H -elim (lexs_inv_next1 … H) -H #Z #K2 #HK12 #H1 #H2 destruct -<(ceq_ext_inv_eq … H1) -Z /2 width=3 by ex2_intro/ -qed-. - -(* Basic_2A1: includes: lreq_inv_zero1 lreq_inv_succ1 *) -lemma lreq_inv_push1: ∀g,J1,K1,Y. K1.ⓘ{J1} ≡[⫯g] Y → - ∃∃J2,K2. K1 ≡[g] K2 & Y = K2.ⓘ{J2}. -#g #J1 #K1 #Y #H elim (lexs_inv_push1 … H) -H /2 width=4 by ex2_2_intro/ -qed-. - -(* Basic_2A1: includes: lreq_inv_atom2 *) -lemma lreq_inv_atom2: ∀f,X. X ≡[f] ⋆ → X = ⋆. -/2 width=4 by lexs_inv_atom2/ qed-. - -(* Basic_2A1: includes: lreq_inv_pair2 *) -lemma lreq_inv_next2: ∀g,J,X,K2. X ≡[↑g] K2.ⓘ{J} → - ∃∃K1. K1 ≡[g] K2 & X = K1.ⓘ{J}. -#g #J #X #K2 #H -elim (lexs_inv_next2 … H) -H #Z #K1 #HK12 #H1 #H2 destruct -<(ceq_ext_inv_eq … H1) -J /2 width=3 by ex2_intro/ -qed-. - -(* Basic_2A1: includes: lreq_inv_zero2 lreq_inv_succ2 *) -lemma lreq_inv_push2: ∀g,J2,X,K2. X ≡[⫯g] K2.ⓘ{J2} → - ∃∃J1,K1. K1 ≡[g] K2 & X = K1.ⓘ{J1}. -#g #J2 #X #K2 #H elim (lexs_inv_push2 … H) -H /2 width=4 by ex2_2_intro/ -qed-. - -(* Basic_2A1: includes: lreq_inv_pair *) -lemma lreq_inv_next: ∀f,I1,I2,L1,L2. L1.ⓘ{I1} ≡[↑f] L2.ⓘ{I2} → - L1 ≡[f] L2 ∧ I1 = I2. -#f #I1 #I2 #L1 #L2 #H elim (lexs_inv_next … H) -H -/3 width=3 by ceq_ext_inv_eq, conj/ -qed-. - -(* Basic_2A1: includes: lreq_inv_succ *) -lemma lreq_inv_push: ∀f,I1,I2,L1,L2. L1.ⓘ{I1} ≡[⫯f] L2.ⓘ{I2} → L1 ≡[f] L2. -#f #I1 #I2 #L1 #L2 #H elim (lexs_inv_push … H) -H /2 width=1 by conj/ -qed-. - -lemma lreq_inv_tl: ∀f,I,L1,L2. L1 ≡[⫱f] L2 → L1.ⓘ{I} ≡[f] L2.ⓘ{I}. -/2 width=1 by lexs_inv_tl/ qed-. - -(* Basic_2A1: removed theorems 5: - lreq_pair_lt lreq_succ_lt lreq_pair_O_Y lreq_O2 lreq_inv_O_Y -*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lreq_length.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lreq_length.ma deleted file mode 100644 index fcb54e71a..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lreq_length.ma +++ /dev/null @@ -1,24 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/relocation/lexs_length.ma". -include "basic_2/relocation/lreq.ma". - -(* RANGED EQUIVALENCE FOR LOCAL ENVIRONMENTS ********************************) - -(* Forward lemmas with length for local environments ************************) - -(* Basic_2A1: includes: lreq_fwd_length *) -lemma lreq_fwd_length: ∀f,L1,L2. L1 ≡[f] L2 → |L1| = |L2|. -/2 width=4 by lexs_fwd_length/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lreq_lreq.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lreq_lreq.ma deleted file mode 100644 index d0c06d0c7..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lreq_lreq.ma +++ /dev/null @@ -1,37 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/syntax/ceq_ext_ceq_ext.ma". -include "basic_2/relocation/lexs_lexs.ma". - -(* RANGED EQUIVALENCE FOR LOCAL ENVIRONMENTS ********************************) - -(* Main properties **********************************************************) - -theorem lreq_trans: ∀f. Transitive … (lreq f). -/3 width=5 by lexs_trans, ceq_ext_trans/ qed-. - -theorem lreq_canc_sn: ∀f. left_cancellable … (lreq f). -/3 width=3 by lexs_canc_sn, lreq_trans, lreq_sym/ qed-. - -theorem lreq_canc_dx: ∀f. right_cancellable … (lreq f). -/3 width=3 by lexs_canc_dx, lreq_trans, lreq_sym/ qed-. - -theorem lreq_join: ∀f1,L1,L2. L1 ≡[f1] L2 → ∀f2. L1 ≡[f2] L2 → - ∀f. f1 ⋓ f2 ≘ f → L1 ≡[f] L2. -/2 width=5 by lexs_join/ qed-. - -theorem lreq_meet: ∀f1,L1,L2. L1 ≡[f1] L2 → ∀f2. L1 ≡[f2] L2 → - ∀f. f1 ⋒ f2 ≘ f → L1 ≡[f] L2. -/2 width=5 by lexs_meet/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/seq.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/seq.ma new file mode 100644 index 000000000..8b833f570 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/seq.ma @@ -0,0 +1,103 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/notation/relations/ideqsn_3.ma". +include "basic_2/syntax/ceq_ext.ma". +include "basic_2/relocation/sex.ma". + +(* SYNTACTIC EQUIVALENCE FOR SELECTED LOCAL ENVIRONMENTS ********************) + +(* Basic_2A1: includes: lreq_atom lreq_zero lreq_pair lreq_succ *) +definition seq: relation3 rtmap lenv lenv ≝ sex ceq_ext cfull. + +interpretation + "syntactic equivalence on selected entries (local environment)" + 'IdEqSn f L1 L2 = (seq f L1 L2). + +(* Basic properties *********************************************************) + +lemma seq_eq_repl_back: ∀L1,L2. eq_repl_back … (λf. L1 ≡[f] L2). +/2 width=3 by sex_eq_repl_back/ qed-. + +lemma seq_eq_repl_fwd: ∀L1,L2. eq_repl_fwd … (λf. L1 ≡[f] L2). +/2 width=3 by sex_eq_repl_fwd/ qed-. + +lemma sle_seq_trans: ∀f2,L1,L2. L1 ≡[f2] L2 → + ∀f1. f1 ⊆ f2 → L1 ≡[f1] L2. +/2 width=3 by sle_sex_trans/ qed-. + +(* Basic_2A1: includes: lreq_refl *) +lemma seq_refl: ∀f. reflexive … (seq f). +/2 width=1 by sex_refl/ qed. + +(* Basic_2A1: includes: lreq_sym *) +lemma seq_sym: ∀f. symmetric … (seq f). +/3 width=2 by sex_sym, cext2_sym/ qed-. + +(* Basic inversion lemmas ***************************************************) + +(* Basic_2A1: includes: lreq_inv_atom1 *) +lemma seq_inv_atom1: ∀f,Y. ⋆ ≡[f] Y → Y = ⋆. +/2 width=4 by sex_inv_atom1/ qed-. + +(* Basic_2A1: includes: lreq_inv_pair1 *) +lemma seq_inv_next1: ∀g,J,K1,Y. K1.ⓘ{J} ≡[↑g] Y → + ∃∃K2. K1 ≡[g] K2 & Y = K2.ⓘ{J}. +#g #J #K1 #Y #H +elim (sex_inv_next1 … H) -H #Z #K2 #HK12 #H1 #H2 destruct +<(ceq_ext_inv_eq … H1) -Z /2 width=3 by ex2_intro/ +qed-. + +(* Basic_2A1: includes: lreq_inv_zero1 lreq_inv_succ1 *) +lemma seq_inv_push1: ∀g,J1,K1,Y. K1.ⓘ{J1} ≡[⫯g] Y → + ∃∃J2,K2. K1 ≡[g] K2 & Y = K2.ⓘ{J2}. +#g #J1 #K1 #Y #H elim (sex_inv_push1 … H) -H /2 width=4 by ex2_2_intro/ +qed-. + +(* Basic_2A1: includes: lreq_inv_atom2 *) +lemma seq_inv_atom2: ∀f,X. X ≡[f] ⋆ → X = ⋆. +/2 width=4 by sex_inv_atom2/ qed-. + +(* Basic_2A1: includes: lreq_inv_pair2 *) +lemma seq_inv_next2: ∀g,J,X,K2. X ≡[↑g] K2.ⓘ{J} → + ∃∃K1. K1 ≡[g] K2 & X = K1.ⓘ{J}. +#g #J #X #K2 #H +elim (sex_inv_next2 … H) -H #Z #K1 #HK12 #H1 #H2 destruct +<(ceq_ext_inv_eq … H1) -J /2 width=3 by ex2_intro/ +qed-. + +(* Basic_2A1: includes: lreq_inv_zero2 lreq_inv_succ2 *) +lemma seq_inv_push2: ∀g,J2,X,K2. X ≡[⫯g] K2.ⓘ{J2} → + ∃∃J1,K1. K1 ≡[g] K2 & X = K1.ⓘ{J1}. +#g #J2 #X #K2 #H elim (sex_inv_push2 … H) -H /2 width=4 by ex2_2_intro/ +qed-. + +(* Basic_2A1: includes: lreq_inv_pair *) +lemma seq_inv_next: ∀f,I1,I2,L1,L2. L1.ⓘ{I1} ≡[↑f] L2.ⓘ{I2} → + ∧∧ L1 ≡[f] L2 & I1 = I2. +#f #I1 #I2 #L1 #L2 #H elim (sex_inv_next … H) -H +/3 width=3 by ceq_ext_inv_eq, conj/ +qed-. + +(* Basic_2A1: includes: lreq_inv_succ *) +lemma seq_inv_push: ∀f,I1,I2,L1,L2. L1.ⓘ{I1} ≡[⫯f] L2.ⓘ{I2} → L1 ≡[f] L2. +#f #I1 #I2 #L1 #L2 #H elim (sex_inv_push … H) -H /2 width=1 by conj/ +qed-. + +lemma seq_inv_tl: ∀f,I,L1,L2. L1 ≡[⫱f] L2 → L1.ⓘ{I} ≡[f] L2.ⓘ{I}. +/2 width=1 by sex_inv_tl/ qed-. + +(* Basic_2A1: removed theorems 5: + lreq_pair_lt lreq_succ_lt lreq_pair_O_Y lreq_O2 lreq_inv_O_Y +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/seq_length.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/seq_length.ma new file mode 100644 index 000000000..4055f6638 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/seq_length.ma @@ -0,0 +1,24 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/sex_length.ma". +include "basic_2/relocation/seq.ma". + +(* SYNTACTIC EQUIVALENCE FOR SELECTED LOCAL ENVIRONMENTS ********************) + +(* Forward lemmas with length for local environments ************************) + +(* Basic_2A1: includes: lreq_fwd_length *) +lemma seq_fwd_length: ∀f,L1,L2. L1 ≡[f] L2 → |L1| = |L2|. +/2 width=4 by sex_fwd_length/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/seq_seq.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/seq_seq.ma new file mode 100644 index 000000000..cd8dbe3ea --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/seq_seq.ma @@ -0,0 +1,37 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/syntax/ceq_ext_ceq_ext.ma". +include "basic_2/relocation/sex_sex.ma". + +(* SYNTACTIC EQUIVALENCE FOR SELECTED LOCAL ENVIRONMENTS ********************) + +(* Main properties **********************************************************) + +theorem seq_trans: ∀f. Transitive … (seq f). +/3 width=5 by sex_trans, ceq_ext_trans/ qed-. + +theorem seq_canc_sn: ∀f. left_cancellable … (seq f). +/3 width=3 by sex_canc_sn, seq_trans, seq_sym/ qed-. + +theorem seq_canc_dx: ∀f. right_cancellable … (seq f). +/3 width=3 by sex_canc_dx, seq_trans, seq_sym/ qed-. + +theorem seq_join: ∀f1,L1,L2. L1 ≡[f1] L2 → ∀f2. L1 ≡[f2] L2 → + ∀f. f1 ⋓ f2 ≘ f → L1 ≡[f] L2. +/2 width=5 by sex_join/ qed-. + +theorem seq_meet: ∀f1,L1,L2. L1 ≡[f1] L2 → ∀f2. L1 ≡[f2] L2 → + ∀f. f1 ⋒ f2 ≘ f → L1 ≡[f] L2. +/2 width=5 by sex_meet/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/sex.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/sex.ma new file mode 100644 index 000000000..96f62f93e --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/sex.ma @@ -0,0 +1,296 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "ground_2/relocation/rtmap_sle.ma". +include "ground_2/relocation/rtmap_sdj.ma". +include "basic_2/notation/relations/relation_5.ma". +include "basic_2/syntax/lenv.ma". + +(* GENERIC ENTRYWISE EXTENSION OF CONTEXT-SENSITIVE REALTIONS FOR TERMS *****) + +inductive sex (RN,RP:relation3 lenv bind bind): rtmap → relation lenv ≝ +| sex_atom: ∀f. sex RN RP f (⋆) (⋆) +| sex_next: ∀f,I1,I2,L1,L2. + sex RN RP f L1 L2 → RN L1 I1 I2 → + sex RN RP (↑f) (L1.ⓘ{I1}) (L2.ⓘ{I2}) +| sex_push: ∀f,I1,I2,L1,L2. + sex RN RP f L1 L2 → RP L1 I1 I2 → + sex RN RP (⫯f) (L1.ⓘ{I1}) (L2.ⓘ{I2}) +. + +interpretation "generic entrywise extension (local environment)" + 'Relation RN RP f L1 L2 = (sex RN RP f L1 L2). + +definition R_pw_confluent2_sex: relation3 lenv bind bind → relation3 lenv bind bind → + relation3 lenv bind bind → relation3 lenv bind bind → + relation3 lenv bind bind → relation3 lenv bind bind → + relation3 rtmap lenv bind ≝ + λR1,R2,RN1,RP1,RN2,RP2,f,L0,I0. + ∀I1. R1 L0 I0 I1 → ∀I2. R2 L0 I0 I2 → + ∀L1. L0 ⪤[RN1, RP1, f] L1 → ∀L2. L0 ⪤[RN2, RP2, f] L2 → + ∃∃I. R2 L1 I1 I & R1 L2 I2 I. + +definition sex_transitive: relation3 lenv bind bind → relation3 lenv bind bind → + relation3 lenv bind bind → + relation3 lenv bind bind → relation3 lenv bind bind → + relation3 rtmap lenv bind ≝ + λR1,R2,R3,RN,RP,f,L1,I1. + ∀I. R1 L1 I1 I → ∀L2. L1 ⪤[RN, RP, f] L2 → + ∀I2. R2 L2 I I2 → R3 L1 I1 I2. + +(* Basic inversion lemmas ***************************************************) + +fact sex_inv_atom1_aux: ∀RN,RP,f,X,Y. X ⪤[RN, RP, f] Y → X = ⋆ → Y = ⋆. +#RN #RP #f #X #Y * -f -X -Y // +#f #I1 #I2 #L1 #L2 #_ #_ #H destruct +qed-. + +(* Basic_2A1: includes lpx_sn_inv_atom1 *) +lemma sex_inv_atom1: ∀RN,RP,f,Y. ⋆ ⪤[RN, RP, f] Y → Y = ⋆. +/2 width=6 by sex_inv_atom1_aux/ qed-. + +fact sex_inv_next1_aux: ∀RN,RP,f,X,Y. X ⪤[RN, RP, f] Y → ∀g,J1,K1. X = K1.ⓘ{J1} → f = ↑g → + ∃∃J2,K2. K1 ⪤[RN, RP, g] K2 & RN K1 J1 J2 & Y = K2.ⓘ{J2}. +#RN #RP #f #X #Y * -f -X -Y +[ #f #g #J1 #K1 #H destruct +| #f #I1 #I2 #L1 #L2 #HL #HI #g #J1 #K1 #H1 #H2 <(injective_next … H2) -g destruct + /2 width=5 by ex3_2_intro/ +| #f #I1 #I2 #L1 #L2 #_ #_ #g #J1 #K1 #_ #H elim (discr_push_next … H) +] +qed-. + +(* Basic_2A1: includes lpx_sn_inv_pair1 *) +lemma sex_inv_next1: ∀RN,RP,g,J1,K1,Y. K1.ⓘ{J1} ⪤[RN, RP, ↑g] Y → + ∃∃J2,K2. K1 ⪤[RN, RP, g] K2 & RN K1 J1 J2 & Y = K2.ⓘ{J2}. +/2 width=7 by sex_inv_next1_aux/ qed-. + +fact sex_inv_push1_aux: ∀RN,RP,f,X,Y. X ⪤[RN, RP, f] Y → ∀g,J1,K1. X = K1.ⓘ{J1} → f = ⫯g → + ∃∃J2,K2. K1 ⪤[RN, RP, g] K2 & RP K1 J1 J2 & Y = K2.ⓘ{J2}. +#RN #RP #f #X #Y * -f -X -Y +[ #f #g #J1 #K1 #H destruct +| #f #I1 #I2 #L1 #L2 #_ #_ #g #J1 #K1 #_ #H elim (discr_next_push … H) +| #f #I1 #I2 #L1 #L2 #HL #HI #g #J1 #K1 #H1 #H2 <(injective_push … H2) -g destruct + /2 width=5 by ex3_2_intro/ +] +qed-. + +lemma sex_inv_push1: ∀RN,RP,g,J1,K1,Y. K1.ⓘ{J1} ⪤[RN, RP, ⫯g] Y → + ∃∃J2,K2. K1 ⪤[RN, RP, g] K2 & RP K1 J1 J2 & Y = K2.ⓘ{J2}. +/2 width=7 by sex_inv_push1_aux/ qed-. + +fact sex_inv_atom2_aux: ∀RN,RP,f,X,Y. X ⪤[RN, RP, f] Y → Y = ⋆ → X = ⋆. +#RN #RP #f #X #Y * -f -X -Y // +#f #I1 #I2 #L1 #L2 #_ #_ #H destruct +qed-. + +(* Basic_2A1: includes lpx_sn_inv_atom2 *) +lemma sex_inv_atom2: ∀RN,RP,f,X. X ⪤[RN, RP, f] ⋆ → X = ⋆. +/2 width=6 by sex_inv_atom2_aux/ qed-. + +fact sex_inv_next2_aux: ∀RN,RP,f,X,Y. X ⪤[RN, RP, f] Y → ∀g,J2,K2. Y = K2.ⓘ{J2} → f = ↑g → + ∃∃J1,K1. K1 ⪤[RN, RP, g] K2 & RN K1 J1 J2 & X = K1.ⓘ{J1}. +#RN #RP #f #X #Y * -f -X -Y +[ #f #g #J2 #K2 #H destruct +| #f #I1 #I2 #L1 #L2 #HL #HI #g #J2 #K2 #H1 #H2 <(injective_next … H2) -g destruct + /2 width=5 by ex3_2_intro/ +| #f #I1 #I2 #L1 #L2 #_ #_ #g #J2 #K2 #_ #H elim (discr_push_next … H) +] +qed-. + +(* Basic_2A1: includes lpx_sn_inv_pair2 *) +lemma sex_inv_next2: ∀RN,RP,g,J2,X,K2. X ⪤[RN, RP, ↑g] K2.ⓘ{J2} → + ∃∃J1,K1. K1 ⪤[RN, RP, g] K2 & RN K1 J1 J2 & X = K1.ⓘ{J1}. +/2 width=7 by sex_inv_next2_aux/ qed-. + +fact sex_inv_push2_aux: ∀RN,RP,f,X,Y. X ⪤[RN, RP, f] Y → ∀g,J2,K2. Y = K2.ⓘ{J2} → f = ⫯g → + ∃∃J1,K1. K1 ⪤[RN, RP, g] K2 & RP K1 J1 J2 & X = K1.ⓘ{J1}. +#RN #RP #f #X #Y * -f -X -Y +[ #f #J2 #K2 #g #H destruct +| #f #I1 #I2 #L1 #L2 #_ #_ #g #J2 #K2 #_ #H elim (discr_next_push … H) +| #f #I1 #I2 #L1 #L2 #HL #HI #g #J2 #K2 #H1 #H2 <(injective_push … H2) -g destruct + /2 width=5 by ex3_2_intro/ +] +qed-. + +lemma sex_inv_push2: ∀RN,RP,g,J2,X,K2. X ⪤[RN, RP, ⫯g] K2.ⓘ{J2} → + ∃∃J1,K1. K1 ⪤[RN, RP, g] K2 & RP K1 J1 J2 & X = K1.ⓘ{J1}. +/2 width=7 by sex_inv_push2_aux/ qed-. + +(* Basic_2A1: includes lpx_sn_inv_pair *) +lemma sex_inv_next: ∀RN,RP,f,I1,I2,L1,L2. + L1.ⓘ{I1} ⪤[RN, RP, ↑f] L2.ⓘ{I2} → + L1 ⪤[RN, RP, f] L2 ∧ RN L1 I1 I2. +#RN #RP #f #I1 #I2 #L1 #L2 #H elim (sex_inv_next1 … H) -H +#I0 #L0 #HL10 #HI10 #H destruct /2 width=1 by conj/ +qed-. + +lemma sex_inv_push: ∀RN,RP,f,I1,I2,L1,L2. + L1.ⓘ{I1} ⪤[RN, RP, ⫯f] L2.ⓘ{I2} → + L1 ⪤[RN, RP, f] L2 ∧ RP L1 I1 I2. +#RN #RP #f #I1 #I2 #L1 #L2 #H elim (sex_inv_push1 … H) -H +#I0 #L0 #HL10 #HI10 #H destruct /2 width=1 by conj/ +qed-. + +lemma sex_inv_tl: ∀RN,RP,f,I1,I2,L1,L2. L1 ⪤[RN, RP, ⫱f] L2 → + RN L1 I1 I2 → RP L1 I1 I2 → + L1.ⓘ{I1} ⪤[RN, RP, f] L2.ⓘ{I2}. +#RN #RP #f #I1 #I2 #L2 #L2 elim (pn_split f) * +/2 width=1 by sex_next, sex_push/ +qed-. + +(* Basic forward lemmas *****************************************************) + +lemma sex_fwd_bind: ∀RN,RP,f,I1,I2,L1,L2. + L1.ⓘ{I1} ⪤[RN, RP, f] L2.ⓘ{I2} → + L1 ⪤[RN, RP, ⫱f] L2. +#RN #RP #f #I1 #I2 #L1 #L2 #Hf +elim (pn_split f) * #g #H destruct +[ elim (sex_inv_push … Hf) | elim (sex_inv_next … Hf) ] -Hf // +qed-. + +(* Basic properties *********************************************************) + +lemma sex_eq_repl_back: ∀RN,RP,L1,L2. eq_repl_back … (λf. L1 ⪤[RN, RP, f] L2). +#RN #RP #L1 #L2 #f1 #H elim H -f1 -L1 -L2 // +#f1 #I1 #I2 #L1 #L2 #_ #HI #IH #f2 #H +[ elim (eq_inv_nx … H) -H /3 width=3 by sex_next/ +| elim (eq_inv_px … H) -H /3 width=3 by sex_push/ +] +qed-. + +lemma sex_eq_repl_fwd: ∀RN,RP,L1,L2. eq_repl_fwd … (λf. L1 ⪤[RN, RP, f] L2). +#RN #RP #L1 #L2 @eq_repl_sym /2 width=3 by sex_eq_repl_back/ (**) (* full auto fails *) +qed-. + +lemma sex_refl: ∀RN,RP. c_reflexive … RN → c_reflexive … RP → + ∀f.reflexive … (sex RN RP f). +#RN #RP #HRN #HRP #f #L generalize in match f; -f elim L -L // +#L #I #IH #f elim (pn_split f) * +#g #H destruct /2 width=1 by sex_next, sex_push/ +qed. + +lemma sex_sym: ∀RN,RP. + (∀L1,L2,I1,I2. RN L1 I1 I2 → RN L2 I2 I1) → + (∀L1,L2,I1,I2. RP L1 I1 I2 → RP L2 I2 I1) → + ∀f. symmetric … (sex RN RP f). +#RN #RP #HRN #HRP #f #L1 #L2 #H elim H -L1 -L2 -f +/3 width=2 by sex_next, sex_push/ +qed-. + +lemma sex_pair_repl: ∀RN,RP,f,I1,I2,L1,L2. + L1.ⓘ{I1} ⪤[RN, RP, f] L2.ⓘ{I2} → + ∀J1,J2. RN L1 J1 J2 → RP L1 J1 J2 → + L1.ⓘ{J1} ⪤[RN, RP, f] L2.ⓘ{J2}. +/3 width=3 by sex_inv_tl, sex_fwd_bind/ qed-. + +lemma sex_co: ∀RN1,RP1,RN2,RP2. RN1 ⊆ RN2 → RP1 ⊆ RP2 → + ∀f,L1,L2. L1 ⪤[RN1, RP1, f] L2 → L1 ⪤[RN2, RP2, f] L2. +#RN1 #RP1 #RN2 #RP2 #HRN #HRP #f #L1 #L2 #H elim H -f -L1 -L2 +/3 width=1 by sex_atom, sex_next, sex_push/ +qed-. + +lemma sex_co_isid: ∀RN1,RP1,RN2,RP2. RP1 ⊆ RP2 → + ∀f,L1,L2. L1 ⪤[RN1, RP1, f] L2 → 𝐈⦃f⦄ → + L1 ⪤[RN2, RP2, f] L2. +#RN1 #RP1 #RN2 #RP2 #HR #f #L1 #L2 #H elim H -f -L1 -L2 // +#f #I1 #I2 #K1 #K2 #_ #HI12 #IH #H +[ elim (isid_inv_next … H) -H // +| /4 width=3 by sex_push, isid_inv_push/ +] +qed-. + +lemma sex_sdj: ∀RN,RP. RP ⊆ RN → + ∀f1,L1,L2. L1 ⪤[RN, RP, f1] L2 → + ∀f2. f1 ∥ f2 → L1 ⪤[RP, RN, f2] L2. +#RN #RP #HR #f1 #L1 #L2 #H elim H -f1 -L1 -L2 // +#f1 #I1 #I2 #L1 #L2 #_ #HI12 #IH #f2 #H12 +[ elim (sdj_inv_nx … H12) -H12 [2,3: // ] + #g2 #H #H2 destruct /3 width=1 by sex_push/ +| elim (sdj_inv_px … H12) -H12 [2,4: // ] * + #g2 #H #H2 destruct /3 width=1 by sex_next, sex_push/ +] +qed-. + +lemma sle_sex_trans: ∀RN,RP. RN ⊆ RP → + ∀f2,L1,L2. L1 ⪤[RN, RP, f2] L2 → + ∀f1. f1 ⊆ f2 → L1 ⪤[RN, RP, f1] L2. +#RN #RP #HR #f2 #L1 #L2 #H elim H -f2 -L1 -L2 // +#f2 #I1 #I2 #L1 #L2 #_ #HI12 #IH #f1 #H12 +[ elim (pn_split f1) * ] +[ /4 width=5 by sex_push, sle_inv_pn/ +| /4 width=5 by sex_next, sle_inv_nn/ +| elim (sle_inv_xp … H12) -H12 [2,3: // ] + #g1 #H #H1 destruct /3 width=5 by sex_push/ +] +qed-. + +lemma sle_sex_conf: ∀RN,RP. RP ⊆ RN → + ∀f1,L1,L2. L1 ⪤[RN, RP, f1] L2 → + ∀f2. f1 ⊆ f2 → L1 ⪤[RN, RP, f2] L2. +#RN #RP #HR #f1 #L1 #L2 #H elim H -f1 -L1 -L2 // +#f1 #I1 #I2 #L1 #L2 #_ #HI12 #IH #f2 #H12 +[2: elim (pn_split f2) * ] +[ /4 width=5 by sex_push, sle_inv_pp/ +| /4 width=5 by sex_next, sle_inv_pn/ +| elim (sle_inv_nx … H12) -H12 [2,3: // ] + #g2 #H #H2 destruct /3 width=5 by sex_next/ +] +qed-. + +lemma sex_sle_split: ∀R1,R2,RP. c_reflexive … R1 → c_reflexive … R2 → + ∀f,L1,L2. L1 ⪤[R1, RP, f] L2 → ∀g. f ⊆ g → + ∃∃L. L1 ⪤[R1, RP, g] L & L ⪤[R2, cfull, f] L2. +#R1 #R2 #RP #HR1 #HR2 #f #L1 #L2 #H elim H -f -L1 -L2 +[ /2 width=3 by sex_atom, ex2_intro/ ] +#f #I1 #I2 #L1 #L2 #_ #HI12 #IH #y #H +[ elim (sle_inv_nx … H ??) -H [ |*: // ] #g #Hfg #H destruct + elim (IH … Hfg) -IH -Hfg /3 width=5 by sex_next, ex2_intro/ +| elim (sle_inv_px … H ??) -H [1,3: * |*: // ] #g #Hfg #H destruct + elim (IH … Hfg) -IH -Hfg /3 width=5 by sex_next, sex_push, ex2_intro/ +] +qed-. + +lemma sex_sdj_split: ∀R1,R2,RP. c_reflexive … R1 → c_reflexive … R2 → + ∀f,L1,L2. L1 ⪤[R1, RP, f] L2 → ∀g. f ∥ g → + ∃∃L. L1 ⪤[RP, R1, g] L & L ⪤[R2, cfull, f] L2. +#R1 #R2 #RP #HR1 #HR2 #f #L1 #L2 #H elim H -f -L1 -L2 +[ /2 width=3 by sex_atom, ex2_intro/ ] +#f #I1 #I2 #L1 #L2 #_ #HI12 #IH #y #H +[ elim (sdj_inv_nx … H ??) -H [ |*: // ] #g #Hfg #H destruct + elim (IH … Hfg) -IH -Hfg /3 width=5 by sex_next, sex_push, ex2_intro/ +| elim (sdj_inv_px … H ??) -H [1,3: * |*: // ] #g #Hfg #H destruct + elim (IH … Hfg) -IH -Hfg /3 width=5 by sex_next, sex_push, ex2_intro/ +] +qed-. + +lemma sex_dec: ∀RN,RP. + (∀L,I1,I2. Decidable (RN L I1 I2)) → + (∀L,I1,I2. Decidable (RP L I1 I2)) → + ∀L1,L2,f. Decidable (L1 ⪤[RN, RP, f] L2). +#RN #RP #HRN #HRP #L1 elim L1 -L1 [ * | #L1 #I1 #IH * ] +[ /2 width=1 by sex_atom, or_introl/ +| #L2 #I2 #f @or_intror #H + lapply (sex_inv_atom1 … H) -H #H destruct +| #f @or_intror #H + lapply (sex_inv_atom2 … H) -H #H destruct +| #L2 #I2 #f elim (IH L2 (⫱f)) -IH #HL12 + [2: /4 width=3 by sex_fwd_bind, or_intror/ ] + elim (pn_split f) * #g #H destruct + [ elim (HRP L1 I1 I2) | elim (HRN L1 I1 I2) ] -HRP -HRN #HV12 + [1,3: /3 width=1 by sex_push, sex_next, or_introl/ ] + @or_intror #H + [ elim (sex_inv_push … H) | elim (sex_inv_next … H) ] -H + /2 width=1 by/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/sex_length.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/sex_length.ma new file mode 100644 index 000000000..86473e631 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/sex_length.ma @@ -0,0 +1,46 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/syntax/lenv_length.ma". +include "basic_2/relocation/sex.ma". + +(* GENERIC ENTRYWISE EXTENSION OF CONTEXT-SENSITIVE REALTIONS FOR TERMS *****) + +(* Forward lemmas with length for local environments ************************) + +lemma sex_fwd_length: ∀RN,RP,f,L1,L2. L1 ⪤[RN, RP, f] L2 → |L1| = |L2|. +#RN #RP #f #L1 #L2 #H elim H -f -L1 -L2 // +#f #I1 #I2 #L1 #L2 >length_bind >length_bind // +qed-. + +(* Properties with length for local environments ****************************) + +lemma sex_length_cfull: ∀L1,L2. |L1| = |L2| → ∀f. L1 ⪤[cfull, cfull, f] L2. +#L1 elim L1 -L1 +[ #Y2 #H >(length_inv_zero_sn … H) -Y2 // +| #L1 #I1 #IH #Y2 #H #f + elim (length_inv_succ_sn … H) -H #I2 #L2 #HL12 #H destruct + elim (pn_split f) * #g #H destruct /3 width=1 by sex_next, sex_push/ +] +qed. + +lemma sex_length_isid: ∀R,L1,L2. |L1| = |L2| → + ∀f. 𝐈⦃f⦄ → L1 ⪤[R, cfull, f] L2. +#R #L1 elim L1 -L1 +[ #Y2 #H >(length_inv_zero_sn … H) -Y2 // +| #L1 #I1 #IH #Y2 #H #f #Hf + elim (length_inv_succ_sn … H) -H #I2 #L2 #HL12 #H destruct + elim (isid_inv_gen … Hf) -Hf #g #Hg #H destruct /3 width=1 by sex_push/ +] +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/sex_sex.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/sex_sex.ma new file mode 100644 index 000000000..df181b7f6 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/sex_sex.ma @@ -0,0 +1,119 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "ground_2/relocation/rtmap_sand.ma". +include "basic_2/relocation/drops.ma". + +(* GENERIC ENTRYWISE EXTENSION OF CONTEXT-SENSITIVE REALTIONS FOR TERMS *****) + +(* Main properties **********************************************************) + +theorem sex_trans_gen (RN1) (RP1) (RN2) (RP2) (RN) (RP): + ∀L1,f. + (∀g,I,K,n. ⬇*[n] L1 ≘ K.ⓘ{I} → ↑g = ⫱*[n] f → sex_transitive RN1 RN2 RN RN1 RP1 g K I) → + (∀g,I,K,n. ⬇*[n] L1 ≘ K.ⓘ{I} → ⫯g = ⫱*[n] f → sex_transitive RP1 RP2 RP RN1 RP1 g K I) → + ∀L0. L1 ⪤[RN1, RP1, f] L0 → + ∀L2. L0 ⪤[RN2, RP2, f] L2 → + L1 ⪤[RN, RP, f] L2. +#RN1 #RP1 #RN2 #RP2 #RN #RP #L1 elim L1 -L1 +[ #f #_ #_ #L0 #H1 #L2 #H2 + lapply (sex_inv_atom1 … H1) -H1 #H destruct + lapply (sex_inv_atom1 … H2) -H2 #H destruct + /2 width=1 by sex_atom/ +| #K1 #I1 #IH #f elim (pn_split f) * #g #H destruct + #HN #HP #L0 #H1 #L2 #H2 + [ elim (sex_inv_push1 … H1) -H1 #I0 #K0 #HK10 #HI10 #H destruct + elim (sex_inv_push1 … H2) -H2 #I2 #K2 #HK02 #HI02 #H destruct + lapply (HP … 0 … HI10 … HK10 … HI02) -HI10 -HI02 /2 width=2 by drops_refl/ #HI12 + lapply (IH … HK10 … HK02) -IH -K0 /3 width=3 by sex_push, drops_drop/ + | elim (sex_inv_next1 … H1) -H1 #I0 #K0 #HK10 #HI10 #H destruct + elim (sex_inv_next1 … H2) -H2 #I2 #K2 #HK02 #HI02 #H destruct + lapply (HN … 0 … HI10 … HK10 … HI02) -HI10 -HI02 /2 width=2 by drops_refl/ #HI12 + lapply (IH … HK10 … HK02) -IH -K0 /3 width=3 by sex_next, drops_drop/ + ] +] +qed-. + +theorem sex_trans (RN) (RP) (f): (∀g,I,K. sex_transitive RN RN RN RN RP g K I) → + (∀g,I,K. sex_transitive RP RP RP RN RP g K I) → + Transitive … (sex RN RP f). +/2 width=9 by sex_trans_gen/ qed-. + +theorem sex_trans_id_cfull: ∀R1,R2,R3,L1,L,f. L1 ⪤[R1, cfull, f] L → 𝐈⦃f⦄ → + ∀L2. L ⪤[R2, cfull, f] L2 → L1 ⪤[R3, cfull, f] L2. +#R1 #R2 #R3 #L1 #L #f #H elim H -L1 -L -f +[ #f #Hf #L2 #H >(sex_inv_atom1 … H) -L2 // ] +#f #I1 #I #K1 #K #HK1 #_ #IH #Hf #L2 #H +[ elim (isid_inv_next … Hf) | lapply (isid_inv_push … Hf ??) ] -Hf [5: |*: // ] #Hf +elim (sex_inv_push1 … H) -H #I2 #K2 #HK2 #_ #H destruct +/3 width=1 by sex_push/ +qed-. + +theorem sex_conf (RN1) (RP1) (RN2) (RP2): + ∀L,f. + (∀g,I,K,n. ⬇*[n] L ≘ K.ⓘ{I} → ↑g = ⫱*[n] f → R_pw_confluent2_sex RN1 RN2 RN1 RP1 RN2 RP2 g K I) → + (∀g,I,K,n. ⬇*[n] L ≘ K.ⓘ{I} → ⫯g = ⫱*[n] f → R_pw_confluent2_sex RP1 RP2 RN1 RP1 RN2 RP2 g K I) → + pw_confluent2 … (sex RN1 RP1 f) (sex RN2 RP2 f) L. +#RN1 #RP1 #RN2 #RP2 #L elim L -L +[ #f #_ #_ #L1 #H1 #L2 #H2 >(sex_inv_atom1 … H1) >(sex_inv_atom1 … H2) -H2 -H1 + /2 width=3 by sex_atom, ex2_intro/ +| #L #I0 #IH #f elim (pn_split f) * #g #H destruct + #HN #HP #Y1 #H1 #Y2 #H2 + [ elim (sex_inv_push1 … H1) -H1 #I1 #L1 #HL1 #HI01 #H destruct + elim (sex_inv_push1 … H2) -H2 #I2 #L2 #HL2 #HI02 #H destruct + elim (HP … 0 … HI01 … HI02 … HL1 … HL2) -HI01 -HI02 /2 width=2 by drops_refl/ #I #HI1 #HI2 + elim (IH … HL1 … HL2) -IH -HL1 -HL2 /3 width=5 by drops_drop, sex_push, ex2_intro/ + | elim (sex_inv_next1 … H1) -H1 #I1 #L1 #HL1 #HI01 #H destruct + elim (sex_inv_next1 … H2) -H2 #I2 #L2 #HL2 #HI02 #H destruct + elim (HN … 0 … HI01 … HI02 … HL1 … HL2) -HI01 -HI02 /2 width=2 by drops_refl/ #I #HI1 #HI2 + elim (IH … HL1 … HL2) -IH -HL1 -HL2 /3 width=5 by drops_drop, sex_next, ex2_intro/ + ] +] +qed-. + +theorem sex_canc_sn: ∀RN,RP,f. Transitive … (sex RN RP f) → + symmetric … (sex RN RP f) → + left_cancellable … (sex RN RP f). +/3 width=3 by/ qed-. + +theorem sex_canc_dx: ∀RN,RP,f. Transitive … (sex RN RP f) → + symmetric … (sex RN RP f) → + right_cancellable … (sex RN RP f). +/3 width=3 by/ qed-. + +lemma sex_meet: ∀RN,RP,L1,L2. + ∀f1. L1 ⪤[RN, RP, f1] L2 → + ∀f2. L1 ⪤[RN, RP, f2] L2 → + ∀f. f1 ⋒ f2 ≘ f → L1 ⪤[RN, RP, f] L2. +#RN #RP #L1 #L2 #f1 #H elim H -f1 -L1 -L2 // +#f1 #I1 #I2 #L1 #L2 #_ #HI12 #IH #f2 #H #f #Hf +elim (pn_split f2) * #g2 #H2 destruct +try elim (sex_inv_push … H) try elim (sex_inv_next … H) -H +[ elim (sand_inv_npx … Hf) | elim (sand_inv_nnx … Hf) +| elim (sand_inv_ppx … Hf) | elim (sand_inv_pnx … Hf) +] -Hf /3 width=5 by sex_next, sex_push/ +qed-. + +lemma sex_join: ∀RN,RP,L1,L2. + ∀f1. L1 ⪤[RN, RP, f1] L2 → + ∀f2. L1 ⪤[RN, RP, f2] L2 → + ∀f. f1 ⋓ f2 ≘ f → L1 ⪤[RN, RP, f] L2. +#RN #RP #L1 #L2 #f1 #H elim H -f1 -L1 -L2 // +#f1 #I1 #I2 #L1 #L2 #_ #HI12 #IH #f2 #H #f #Hf +elim (pn_split f2) * #g2 #H2 destruct +try elim (sex_inv_push … H) try elim (sex_inv_next … H) -H +[ elim (sor_inv_npx … Hf) | elim (sor_inv_nnx … Hf) +| elim (sor_inv_ppx … Hf) | elim (sor_inv_pnx … Hf) +] -Hf /3 width=5 by sex_next, sex_push/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/sex_tc.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/sex_tc.ma new file mode 100644 index 000000000..b2134410b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/sex_tc.ma @@ -0,0 +1,119 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "ground_2/lib/star.ma". +include "basic_2/relocation/sex.ma". + +(* GENERIC ENTRYWISE EXTENSION OF CONTEXT-SENSITIVE REALTIONS FOR TERMS *****) + +definition s_rs_transitive_isid: relation (relation3 lenv bind bind) ≝ λRN,RP. + ∀f. 𝐈⦃f⦄ → s_rs_transitive … RP (λ_.sex RN RP f). + +(* Properties with transitive closure ***************************************) + +lemma sex_tc_refl: ∀RN,RP. c_reflexive … RN → c_reflexive … RP → + ∀f. reflexive … (TC … (sex RN RP f)). +/3 width=1 by sex_refl, TC_reflexive/ qed. + +lemma sex_tc_next_sn: ∀RN,RP. c_reflexive … RN → + ∀f,I2,L1,L2. TC … (sex RN RP f) L1 L2 → ∀I1. RN L1 I1 I2 → + TC … (sex RN RP (↑f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}). +#RN #RP #HRN #f #I2 #L1 #L2 #H @(TC_ind_dx ??????? H) -L1 +/3 width=3 by sex_next, TC_strap, inj/ +qed. + +lemma sex_tc_next_dx: ∀RN,RP. c_reflexive … RN → c_reflexive … RP → + ∀f,I1,I2,L1. (CTC … RN) L1 I1 I2 → ∀L2. L1 ⪤[RN, RP, f] L2 → + TC … (sex RN RP (↑f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}). +#RN #RP #HRN #HRP #f #I1 #I2 #L1 #H elim H -I2 +/4 width=5 by sex_refl, sex_next, step, inj/ +qed. + +lemma sex_tc_push_sn: ∀RN,RP. c_reflexive … RP → + ∀f,I2,L1,L2. TC … (sex RN RP f) L1 L2 → ∀I1. RP L1 I1 I2 → + TC … (sex RN RP (⫯f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}). +#RN #RP #HRP #f #I2 #L1 #L2 #H @(TC_ind_dx ??????? H) -L1 +/3 width=3 by sex_push, TC_strap, inj/ +qed. + +lemma sex_tc_push_dx: ∀RN,RP. c_reflexive … RN → c_reflexive … RP → + ∀f,I1,I2,L1. (CTC … RP) L1 I1 I2 → ∀L2. L1 ⪤[RN, RP, f] L2 → + TC … (sex RN RP (⫯f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}). +#RN #RP #HRN #HRP #f #I1 #I2 #L1 #H elim H -I2 +/4 width=5 by sex_refl, sex_push, step, inj/ +qed. + +lemma sex_tc_inj_sn: ∀RN,RP,f,L1,L2. L1 ⪤[RN, RP, f] L2 → L1 ⪤[CTC … RN, RP, f] L2. +#RN #RP #f #L1 #L2 #H elim H -f -L1 -L2 +/3 width=1 by sex_push, sex_next, inj/ +qed. + +lemma sex_tc_inj_dx: ∀RN,RP,f,L1,L2. L1 ⪤[RN, RP, f] L2 → L1 ⪤[RN, CTC … RP, f] L2. +#RN #RP #f #L1 #L2 #H elim H -f -L1 -L2 +/3 width=1 by sex_push, sex_next, inj/ +qed. + +(* Main properties with transitive closure **********************************) + +theorem sex_tc_next: ∀RN,RP. c_reflexive … RN → c_reflexive … RP → + ∀f,I1,I2,L1. (CTC … RN) L1 I1 I2 → ∀L2. TC … (sex RN RP f) L1 L2 → + TC … (sex RN RP (↑f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}). +#RN #RP #HRN #HRP #f #I1 #I2 #L1 #H elim H -I2 +/4 width=5 by sex_tc_next_sn, sex_tc_refl, trans_TC/ +qed. + +theorem sex_tc_push: ∀RN,RP. c_reflexive … RN → c_reflexive … RP → + ∀f,I1,I2,L1. (CTC … RP) L1 I1 I2 → ∀L2. TC … (sex RN RP f) L1 L2 → + TC … (sex RN RP (⫯f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}). +#RN #RP #HRN #HRP #f #I1 #I2 #L1 #H elim H -I2 +/4 width=5 by sex_tc_push_sn, sex_tc_refl, trans_TC/ +qed. + +(* Basic_2A1: uses: TC_lpx_sn_ind *) +theorem sex_tc_step_dx: ∀RN,RP. s_rs_transitive_isid RN RP → + ∀f,L1,L. L1 ⪤[RN, RP, f] L → 𝐈⦃f⦄ → + ∀L2. L ⪤[RN, CTC … RP, f] L2 → L1⪤ [RN, CTC … RP, f] L2. +#RN #RP #HRP #f #L1 #L #H elim H -f -L1 -L +[ #f #_ #Y #H -HRP >(sex_inv_atom1 … H) -Y // ] +#f #I1 #I #L1 #L #HL1 #HI1 #IH #Hf #Y #H +[ elim (isid_inv_next … Hf) -Hf // +| lapply (isid_inv_push … Hf ??) -Hf [3: |*: // ] #Hf + elim (sex_inv_push1 … H) -H #I2 #L2 #HL2 #HI2 #H destruct + @sex_push [ /2 width=1 by/ ] -L2 -IH + @(TC_strap … HI1) -HI1 + @(HRP … HL1) // (**) (* auto fails *) +] +qed-. + +(* Advanced properties ******************************************************) + +lemma sex_tc_dx: ∀RN,RP. s_rs_transitive_isid RN RP → + ∀f. 𝐈⦃f⦄ → ∀L1,L2. TC … (sex RN RP f) L1 L2 → L1 ⪤[RN, CTC … RP, f] L2. +#RN #RP #HRP #f #Hf #L1 #L2 #H @(TC_ind_dx ??????? H) -L1 +/3 width=3 by sex_tc_step_dx, sex_tc_inj_dx/ +qed. + +(* Advanced inversion lemmas ************************************************) + +lemma sex_inv_tc_sn: ∀RN,RP. c_reflexive … RN → c_reflexive … RP → + ∀f,L1,L2. L1 ⪤[CTC … RN, RP, f] L2 → TC … (sex RN RP f) L1 L2. +#RN #RP #HRN #HRP #f #L1 #L2 #H elim H -f -L1 -L2 +/2 width=1 by sex_tc_next, sex_tc_push_sn, sex_atom, inj/ +qed-. + +lemma sex_inv_tc_dx: ∀RN,RP. c_reflexive … RN → c_reflexive … RP → + ∀f,L1,L2. L1 ⪤[RN, CTC … RP, f] L2 → TC … (sex RN RP f) L1 L2. +#RN #RP #HRN #HRP #f #L1 #L2 #H elim H -f -L1 -L2 +/2 width=1 by sex_tc_push, sex_tc_next_sn, sex_atom, inj/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_fdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_fdeq.ma new file mode 100644 index 000000000..16675efdb --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_fdeq.ma @@ -0,0 +1,31 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/fdeq.ma". +include "basic_2/rt_computation/cpxs_rdeq.ma". + +(* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS **************) + +(* Properties with degree-based equivalence for closures ********************) + +lemma fdeq_cpxs_trans: ∀h,o,G1,G2,L1,L2,T1,T. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T⦄ → + ∀T2. ⦃G2, L2⦄ ⊢ T ⬈*[h] T2 → + ∃∃T0. ⦃G1, L1⦄ ⊢ T1 ⬈*[h] T0 & ⦃G1, L1, T0⦄ ≛[h, o] ⦃G2, L2, T2⦄. +#h #o #G1 #G2 #L1 #L2 #T1 #T #H #T2 #HT2 +elim (fdeq_inv_gen_dx … H) -H #H #HL12 #HT1 destruct +elim (rdeq_cpxs_trans … HT2 … HL12) #T0 #HT0 #HT02 +lapply (cpxs_rdeq_conf_dx … HT2 … HL12) -HL12 #HL12 +elim (tdeq_cpxs_trans … HT1 … HT0) -T #T #HT1 #HT0 +/4 width=5 by fdeq_intro_dx, tdeq_trans, ex2_intro/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_ffdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_ffdeq.ma deleted file mode 100644 index c0b475a3e..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_ffdeq.ma +++ /dev/null @@ -1,31 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/ffdeq.ma". -include "basic_2/rt_computation/cpxs_lfdeq.ma". - -(* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS **************) - -(* Properties with degree-based equivalence for closures ********************) - -lemma ffdeq_cpxs_trans: ∀h,o,G1,G2,L1,L2,T1,T. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T⦄ → - ∀T2. ⦃G2, L2⦄ ⊢ T ⬈*[h] T2 → - ∃∃T0. ⦃G1, L1⦄ ⊢ T1 ⬈*[h] T0 & ⦃G1, L1, T0⦄ ≛[h, o] ⦃G2, L2, T2⦄. -#h #o #G1 #G2 #L1 #L2 #T1 #T #H #T2 #HT2 -elim (ffdeq_inv_gen_dx … H) -H #H #HL12 #HT1 destruct -elim (lfdeq_cpxs_trans … HT2 … HL12) #T0 #HT0 #HT02 -lapply (cpxs_lfdeq_conf_dx … HT2 … HL12) -HL12 #HL12 -elim (tdeq_cpxs_trans … HT1 … HT0) -T #T #HT1 #HT0 -/4 width=5 by ffdeq_intro_dx, tdeq_trans, ex2_intro/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_lfdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_lfdeq.ma deleted file mode 100644 index d8cfade21..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_lfdeq.ma +++ /dev/null @@ -1,49 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/rt_transition/cpx_lfdeq.ma". -include "basic_2/rt_computation/cpxs_tdeq.ma". - -(* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS **************) - -(* Properties with degree-based equivalence for local environments **********) - -(* Basic_2A1: was just: lleq_cpxs_trans *) -lemma lfdeq_cpxs_trans: ∀h,o,G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ⬈*[h] T1 → - ∀L2. L2 ≛[h, o, T0] L0 → - ∃∃T. ⦃G, L2⦄ ⊢ T0 ⬈*[h] T & T ≛[h, o] T1. -#h #o #G #L0 #T0 #T1 #H @(cpxs_ind_dx … H) -T0 /2 width=3 by ex2_intro/ -#T0 #T #HT0 #_ #IH #L2 #HL2 -elim (lfdeq_cpx_trans … HL2 … HT0) #U1 #H1 #H2 -elim (IH L2) -IH /2 width=4 by cpx_lfdeq_conf_dx/ -L0 #U2 #H3 #H4 -elim (tdeq_cpxs_trans … H2 … H3) -T #U0 #H2 #H3 -/3 width=5 by cpxs_strap2, tdeq_trans, ex2_intro/ -qed-. - -(* Basic_2A1: was just: cpxs_lleq_conf *) -lemma cpxs_lfdeq_conf: ∀h,o,G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ⬈*[h] T1 → - ∀L2. L0 ≛[h, o, T0] L2 → - ∃∃T. ⦃G, L2⦄ ⊢ T0 ⬈*[h] T & T ≛[h, o] T1. -/3 width=3 by lfdeq_cpxs_trans, lfdeq_sym/ qed-. - -(* Basic_2A1: was just: cpxs_lleq_conf_dx *) -lemma cpxs_lfdeq_conf_dx: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ⬈*[h] T2 → - ∀L1. L1 ≛[h, o, T1] L2 → L1 ≛[h, o, T2] L2. -#h #o #G #L2 #T1 #T2 #H @(cpxs_ind … H) -T2 /3 width=6 by cpx_lfdeq_conf_dx/ -qed-. - -(* Basic_2A1: was just: lleq_conf_sn *) -lemma cpxs_lfdeq_conf_sn: ∀h,o,G,L1,T1,T2. ⦃G, L1⦄ ⊢ T1 ⬈*[h] T2 → - ∀L2. L1 ≛[h, o, T1] L2 → L1 ≛[h, o, T2] L2. -/4 width=6 by cpxs_lfdeq_conf_dx, lfdeq_sym/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_rdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_rdeq.ma new file mode 100644 index 000000000..292dc54e1 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_rdeq.ma @@ -0,0 +1,49 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/rt_transition/cpx_rdeq.ma". +include "basic_2/rt_computation/cpxs_tdeq.ma". + +(* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS **************) + +(* Properties with degree-based equivalence for local environments **********) + +(* Basic_2A1: was just: lleq_cpxs_trans *) +lemma rdeq_cpxs_trans: ∀h,o,G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ⬈*[h] T1 → + ∀L2. L2 ≛[h, o, T0] L0 → + ∃∃T. ⦃G, L2⦄ ⊢ T0 ⬈*[h] T & T ≛[h, o] T1. +#h #o #G #L0 #T0 #T1 #H @(cpxs_ind_dx … H) -T0 /2 width=3 by ex2_intro/ +#T0 #T #HT0 #_ #IH #L2 #HL2 +elim (rdeq_cpx_trans … HL2 … HT0) #U1 #H1 #H2 +elim (IH L2) -IH /2 width=4 by cpx_rdeq_conf_dx/ -L0 #U2 #H3 #H4 +elim (tdeq_cpxs_trans … H2 … H3) -T #U0 #H2 #H3 +/3 width=5 by cpxs_strap2, tdeq_trans, ex2_intro/ +qed-. + +(* Basic_2A1: was just: cpxs_lleq_conf *) +lemma cpxs_rdeq_conf: ∀h,o,G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ⬈*[h] T1 → + ∀L2. L0 ≛[h, o, T0] L2 → + ∃∃T. ⦃G, L2⦄ ⊢ T0 ⬈*[h] T & T ≛[h, o] T1. +/3 width=3 by rdeq_cpxs_trans, rdeq_sym/ qed-. + +(* Basic_2A1: was just: cpxs_lleq_conf_dx *) +lemma cpxs_rdeq_conf_dx: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ⬈*[h] T2 → + ∀L1. L1 ≛[h, o, T1] L2 → L1 ≛[h, o, T2] L2. +#h #o #G #L2 #T1 #T2 #H @(cpxs_ind … H) -T2 /3 width=6 by cpx_rdeq_conf_dx/ +qed-. + +(* Basic_2A1: was just: lleq_conf_sn *) +lemma cpxs_rdeq_conf_sn: ∀h,o,G,L1,T1,T2. ⦃G, L1⦄ ⊢ T1 ⬈*[h] T2 → + ∀L2. L1 ≛[h, o, T1] L2 → L1 ≛[h, o, T2] L2. +/4 width=6 by cpxs_rdeq_conf_dx, rdeq_sym/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tdeq.ma index f86f4396f..279ee212e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tdeq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tdeq.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/rt_transition/lfpx_lfdeq.ma". +include "basic_2/rt_transition/rpx_rdeq.ma". include "basic_2/rt_computation/cpxs.ma". (* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS **************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx.ma index 90480a605..5efee8db4 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/rt_transition/lpx_lfdeq.ma". +include "basic_2/rt_transition/lpx_rdeq.ma". include "basic_2/rt_computation/csx_drops.ma". (* STRONGLY NORMALIZING TERMS FOR UNBOUND PARALLEL RT-TRANSITION ************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_fdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_fdeq.ma new file mode 100644 index 000000000..f8978c4cd --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_fdeq.ma @@ -0,0 +1,26 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/fdeq.ma". +include "basic_2/rt_computation/csx_rdeq.ma". + +(* STRONGLY NORMALIZING TERMS FOR UNBOUND PARALLEL RT-TRANSITION ************) + +(* Properties with degree-based equivalence for closures ********************) + +lemma csx_fdeq_conf: ∀h,o,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬈*[h, o] 𝐒⦃T1⦄ → + ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G2, L2⦄ ⊢ ⬈*[h, o] 𝐒⦃T2⦄. +#h #o #G1 #L1 #T1 #HT1 #G2 #L2 #T2 * -G2 -L2 -T2 +/3 width=3 by csx_rdeq_conf, csx_tdeq_trans/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_ffdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_ffdeq.ma deleted file mode 100644 index 69a6aa002..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_ffdeq.ma +++ /dev/null @@ -1,26 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/ffdeq.ma". -include "basic_2/rt_computation/csx_lfdeq.ma". - -(* STRONGLY NORMALIZING TERMS FOR UNBOUND PARALLEL RT-TRANSITION ************) - -(* Properties with degree-based equivalence for closures ********************) - -lemma csx_ffdeq_conf: ∀h,o,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬈*[h, o] 𝐒⦃T1⦄ → - ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G2, L2⦄ ⊢ ⬈*[h, o] 𝐒⦃T2⦄. -#h #o #G1 #L1 #T1 #HT1 #G2 #L2 #T2 * -G2 -L2 -T2 -/3 width=3 by csx_lfdeq_conf, csx_tdeq_trans/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_fpbq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_fpbq.ma index 022a4d381..fbc312d03 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_fpbq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_fpbq.ma @@ -14,7 +14,7 @@ include "basic_2/rt_transition/fpbq.ma". include "basic_2/rt_computation/csx_fqus.ma". -include "basic_2/rt_computation/csx_ffdeq.ma". +include "basic_2/rt_computation/csx_fdeq.ma". include "basic_2/rt_computation/csx_lpx.ma". (* STRONGLY NORMALIZING TERMS FOR UNBOUND PARALLEL RT-TRANSITION ************) @@ -25,5 +25,5 @@ include "basic_2/rt_computation/csx_lpx.ma". lemma csx_fpbq_conf: ∀h,o,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬈*[h, o] 𝐒⦃T1⦄ → ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ → ⦃G2, L2⦄ ⊢ ⬈*[h, o] 𝐒⦃T2⦄. #h #o #G1 #L1 #T1 #HT1 #G2 #L2 #T2 * -/2 width=6 by csx_cpx_trans, csx_fquq_conf, csx_lpx_conf, csx_ffdeq_conf/ +/2 width=6 by csx_cpx_trans, csx_fquq_conf, csx_lpx_conf, csx_fdeq_conf/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lfdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lfdeq.ma deleted file mode 100644 index 69f78e094..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lfdeq.ma +++ /dev/null @@ -1,35 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/rt_transition/cpx_lfdeq.ma". -include "basic_2/rt_computation/csx_csx.ma". - -(* STRONGLY NORMALIZING TERMS FOR UNBOUND PARALLEL RT-TRANSITION ************) - -(* Properties with degree-based equivalence for local environments **********) - -(* Basic_2A1: uses: csx_lleq_conf *) -lemma csx_lfdeq_conf: ∀h,o,G,L1,T. ⦃G, L1⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄ → - ∀L2. L1 ≛[h, o, T] L2 → ⦃G, L2⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄. -#h #o #G #L1 #T #H -@(csx_ind … H) -T #T1 #_ #IH #L2 #HL12 -@csx_intro #T2 #HT12 #HnT12 -elim (lfdeq_cpx_trans … HL12 … HT12) -HT12 -/5 width=4 by cpx_lfdeq_conf_sn, csx_tdeq_trans, tdeq_trans/ -qed-. - -(* Basic_2A1: uses: csx_lleq_conf *) -lemma csx_lfdeq_trans: ∀h,o,L1,L2,T. L1 ≛[h, o, T] L2 → - ∀G. ⦃G, L2⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄ → ⦃G, L1⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄. -/3 width=3 by csx_lfdeq_conf, lfdeq_sym/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_rdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_rdeq.ma new file mode 100644 index 000000000..de6114d2f --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_rdeq.ma @@ -0,0 +1,35 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/rt_transition/cpx_rdeq.ma". +include "basic_2/rt_computation/csx_csx.ma". + +(* STRONGLY NORMALIZING TERMS FOR UNBOUND PARALLEL RT-TRANSITION ************) + +(* Properties with degree-based equivalence for local environments **********) + +(* Basic_2A1: uses: csx_lleq_conf *) +lemma csx_rdeq_conf: ∀h,o,G,L1,T. ⦃G, L1⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄ → + ∀L2. L1 ≛[h, o, T] L2 → ⦃G, L2⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄. +#h #o #G #L1 #T #H +@(csx_ind … H) -T #T1 #_ #IH #L2 #HL12 +@csx_intro #T2 #HT12 #HnT12 +elim (rdeq_cpx_trans … HL12 … HT12) -HT12 +/5 width=4 by cpx_rdeq_conf_sn, csx_tdeq_trans, tdeq_trans/ +qed-. + +(* Basic_2A1: uses: csx_lleq_conf *) +lemma csx_rdeq_trans: ∀h,o,L1,L2,T. L1 ≛[h, o, T] L2 → + ∀G. ⦃G, L2⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄ → ⦃G, L1⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄. +/3 width=3 by csx_rdeq_conf, rdeq_sym/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg.ma index de77a268f..e3232f7e1 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg.ma @@ -45,6 +45,6 @@ lemma fpbg_fpbs_trans: ∀h,o,G,G2,L,L2,T,T2. ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, qed-. (* Basic_2A1: uses: fpbg_fleq_trans *) -lemma fpbg_ffdeq_trans: ∀h,o,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ >[h, o] ⦃G, L, T⦄ → - ∀G2,L2,T2. ⦃G, L, T⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >[h, o] ⦃G2, L2, T2⦄. -/3 width=5 by fpbg_fpbq_trans, fpbq_ffdeq/ qed-. +lemma fpbg_fdeq_trans: ∀h,o,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ >[h, o] ⦃G, L, T⦄ → + ∀G2,L2,T2. ⦃G, L, T⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >[h, o] ⦃G2, L2, T2⦄. +/3 width=5 by fpbg_fpbq_trans, fpbq_fdeq/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg_fpbs.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg_fpbs.ma index b188e7880..5608b1ff8 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg_fpbs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg_fpbs.ma @@ -12,8 +12,8 @@ (* *) (**************************************************************************) -include "basic_2/static/ffdeq_fqup.ma". -include "basic_2/static/ffdeq_ffdeq.ma". +include "basic_2/static/fdeq_fqup.ma". +include "basic_2/static/fdeq_fdeq.ma". include "basic_2/rt_transition/fpbq_fpb.ma". include "basic_2/rt_computation/fpbg.ma". @@ -30,11 +30,11 @@ qed-. (* Advanced properties with degree-based equivalence on closures ************) (* Basic_2A1: uses: fleq_fpbg_trans *) -lemma ffdeq_fpbg_trans: ∀h,o,G,G2,L,L2,T,T2. ⦃G, L, T⦄ >[h, o] ⦃G2, L2, T2⦄ → - ∀G1,L1,T1. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ >[h, o] ⦃G2, L2, T2⦄. +lemma fdeq_fpbg_trans: ∀h,o,G,G2,L,L2,T,T2. ⦃G, L, T⦄ >[h, o] ⦃G2, L2, T2⦄ → + ∀G1,L1,T1. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ >[h, o] ⦃G2, L2, T2⦄. #h #o #G #G2 #L #L2 #T #T2 * #G0 #L0 #T0 #H0 #H02 #G1 #L1 #T1 #H1 -elim (ffdeq_fpb_trans … H1 … H0) -G -L -T -/4 width=9 by fpbs_strap2, fpbq_ffdeq, ex2_3_intro/ +elim (fdeq_fpb_trans … H1 … H0) -G -L -T +/4 width=9 by fpbs_strap2, fpbq_fdeq, ex2_3_intro/ qed-. (* Properties with parallel proper rst-reduction on closures ****************) @@ -51,7 +51,7 @@ lemma fpbq_fpbg_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ >[h, o] ⦃G2, L2, T2⦄. #h #o #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 elim (fpbq_inv_fpb … H1) -H1 -/2 width=5 by ffdeq_fpbg_trans, fpb_fpbg_trans/ +/2 width=5 by fdeq_fpbg_trans, fpb_fpbg_trans/ qed-. (* Properties with parallel rst-compuutation on closures ********************) @@ -71,10 +71,10 @@ lemma fpbs_inv_fpbg: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, [ /2 width=1 by or_introl/ | #G #G2 #L #L2 #T #T2 #_ #H2 * #H1 elim (fpbq_inv_fpb … H2) -H2 #H2 - [ /3 width=5 by ffdeq_trans, or_introl/ - | elim (ffdeq_fpb_trans … H1 … H2) -G -L -T - /4 width=5 by ex2_3_intro, or_intror, ffdeq_fpbs/ - | /3 width=5 by fpbg_ffdeq_trans, or_intror/ + [ /3 width=5 by fdeq_trans, or_introl/ + | elim (fdeq_fpb_trans … H1 … H2) -G -L -T + /4 width=5 by ex2_3_intro, or_intror, fdeq_fpbs/ + | /3 width=5 by fpbg_fdeq_trans, or_intror/ | /4 width=5 by fpbg_fpbq_trans, fpb_fpbq, or_intror/ ] ] @@ -86,8 +86,8 @@ lemma fpbs_fpb_trans: ∀h,o,F1,F2,K1,K2,T1,T2. ⦃F1, K1, T1⦄ ≥[h, o] ⦃F2 ∀G2,L2,U2. ⦃F2, K2, T2⦄ ≻[h, o] ⦃G2, L2, U2⦄ → ∃∃G1,L1,U1. ⦃F1, K1, T1⦄ ≻[h, o] ⦃G1, L1, U1⦄ & ⦃G1, L1, U1⦄ ≥[h, o] ⦃G2, L2, U2⦄. #h #o #F1 #F2 #K1 #K2 #T1 #T2 #H elim (fpbs_inv_fpbg … H) -H -[ #H12 #G2 #L2 #U2 #H2 elim (ffdeq_fpb_trans … H12 … H2) -F2 -K2 -T2 - /3 width=5 by ffdeq_fpbs, ex2_3_intro/ +[ #H12 #G2 #L2 #U2 #H2 elim (fdeq_fpb_trans … H12 … H2) -F2 -K2 -T2 + /3 width=5 by fdeq_fpbs, ex2_3_intro/ | * #H1 #H2 #H3 #H4 #H5 #H6 #H7 #H8 #H9 @(ex2_3_intro … H4) -H4 /3 width=5 by fpbs_strap1, fpb_fpbq/ ] diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg_lpxs.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg_lpxs.ma index b33d67dea..71dcdc50a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg_lpxs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg_lpxs.ma @@ -20,9 +20,9 @@ include "basic_2/rt_computation/fpbg.ma". (* Properties with unbound rt-computation on full local environments ********) (* Basic_2A1: uses: lpxs_fpbg *) -lemma lpxs_lfdneq_fpbg: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ⬈*[h] L2 → - (L1 ≛[h, o, T] L2 → ⊥) → ⦃G, L1, T⦄ >[h, o] ⦃G, L2, T⦄. +lemma lpxs_rdneq_fpbg: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ⬈*[h] L2 → + (L1 ≛[h, o, T] L2 → ⊥) → ⦃G, L1, T⦄ >[h, o] ⦃G, L2, T⦄. #h #o #G #L1 #L2 #T #H #H0 -elim (lpxs_lfdneq_inv_step_sn … H … H0) -H -H0 -/4 width=7 by fpb_lpx, lpxs_ffdeq_fpbs, ffdeq_intro_sn, ex2_3_intro/ +elim (lpxs_rdneq_inv_step_sn … H … H0) -H -H0 +/4 width=7 by fpb_lpx, lpxs_fdeq_fpbs, fdeq_intro_sn, ex2_3_intro/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs.ma index af6ca54c9..f5c9176b0 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs.ma @@ -54,22 +54,22 @@ lemma fpbs_strap2: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≽[h, o] /2 width=5 by tri_TC_strap/ qed-. (* Basic_2A1: uses: lleq_fpbs fleq_fpbs *) -lemma ffdeq_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄. -/3 width=1 by fpbq_fpbs, fpbq_ffdeq/ qed. +lemma fdeq_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄. +/3 width=1 by fpbq_fpbs, fpbq_fdeq/ qed. (* Basic_2A1: uses: fpbs_lleq_trans *) -lemma fpbs_ffdeq_trans: ∀h,o,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ → - ∀G2,L2,T2. ⦃G, L, T⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄. -/3 width=9 by fpbs_strap1, fpbq_ffdeq/ qed-. +lemma fpbs_fdeq_trans: ∀h,o,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ → + ∀G2,L2,T2. ⦃G, L, T⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄. +/3 width=9 by fpbs_strap1, fpbq_fdeq/ qed-. (* Basic_2A1: uses: lleq_fpbs_trans *) -lemma ffdeq_fpbs_trans: ∀h,o,G,G2,L,L2,T,T2. ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ → - ∀G1,L1,T1. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄. -/3 width=5 by fpbs_strap2, fpbq_ffdeq/ qed-. +lemma fdeq_fpbs_trans: ∀h,o,G,G2,L,L2,T,T2. ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ → + ∀G1,L1,T1. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄. +/3 width=5 by fpbs_strap2, fpbq_fdeq/ qed-. -lemma tdeq_lfdeq_lpx_fpbs: ∀h,o,T1,T2. T1 ≛[h, o] T2 → ∀L1,L0. L1 ≛[h, o, T2] L0 → - ∀G,L2. ⦃G, L0⦄ ⊢ ⬈[h] L2 → ⦃G, L1, T1⦄ ≥[h, o] ⦃G, L2, T2⦄. -/4 width=5 by ffdeq_fpbs, fpbs_strap1, fpbq_lpx, ffdeq_intro_dx/ qed. +lemma tdeq_rdeq_lpx_fpbs: ∀h,o,T1,T2. T1 ≛[h, o] T2 → ∀L1,L0. L1 ≛[h, o, T2] L0 → + ∀G,L2. ⦃G, L0⦄ ⊢ ⬈[h] L2 → ⦃G, L1, T1⦄ ≥[h, o] ⦃G, L2, T2⦄. +/4 width=5 by fdeq_fpbs, fpbs_strap1, fpbq_lpx, fdeq_intro_dx/ qed. (* Basic_2A1: removed theorems 3: fpb_fpbsa_trans fpbs_fpbsa fpbsa_inv_fpbs diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_cpx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_cpx.ma index 3acf1a168..c3cdaf038 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_cpx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_cpx.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/rt_transition/cpx_ffdeq.ma". +include "basic_2/rt_transition/cpx_fdeq.ma". include "basic_2/rt_computation/lpxs_cpxs.ma". include "basic_2/rt_computation/fpbs_lpxs.ma". @@ -26,8 +26,8 @@ lemma fpbs_cpx_tdneq_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈[h] U1 & T1 ≛[h, o] U1 → ⊥ & ⦃G1, L1, U1⦄ ≥[h, o] ⦃G2, L2, U2⦄. #h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 #HnTU2 elim (fpbs_inv_star … H) -H #G0 #L0 #L3 #T0 #T3 #HT10 #H10 #HL03 #H32 -elim (ffdeq_cpx_trans … H32 … HTU2) -HTU2 #T4 #HT34 #H42 -lapply (ffdeq_tdneq_repl_dx … H32 … H42 … HnTU2) -T2 #HnT34 +elim (fdeq_cpx_trans … H32 … HTU2) -HTU2 #T4 #HT34 #H42 +lapply (fdeq_tdneq_repl_dx … H32 … H42 … HnTU2) -T2 #HnT34 lapply (lpxs_cpx_trans … HT34 … HL03) -HT34 #HT34 elim (fqus_cpxs_trans_tdneq … H10 … HT34 HnT34) -T3 #T2 #HT02 #HnT02 #H24 elim (tdeq_dec h o T1 T0) [ #H10 | -HnT02 #HnT10 ] diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_cpxs.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_cpxs.ma index 3551e1f7e..4f23c0d00 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_cpxs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_cpxs.ma @@ -43,7 +43,7 @@ lemma cpxs_tdeq_fpbs_trans: ∀h,o,G1,L1,T1,T. ⦃G1, L1⦄ ⊢ T1 ⬈*[h] T → lemma cpxs_tdeq_fpbs: ∀h,o,G,L,T1,T. ⦃G, L⦄ ⊢ T1 ⬈*[h] T → ∀T2. T ≛[h, o] T2 → ⦃G, L, T1⦄ ≥[h, o] ⦃G, L, T2⦄. -/4 width=3 by cpxs_fpbs_trans, ffdeq_fpbs, tdeq_ffdeq/ qed. +/4 width=3 by cpxs_fpbs_trans, fdeq_fpbs, tdeq_fdeq/ qed. (* Properties with star-iterated structural successor for closures **********) diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_fqup.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_fqup.ma index ce248aeb1..b1e12a7a6 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_fqup.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_fqup.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "basic_2/s_computation/fqus_fqup.ma". -include "basic_2/static/ffdeq_fqup.ma". +include "basic_2/static/fdeq_fqup.ma". include "basic_2/rt_computation/fpbs_fqus.ma". (* PARALLEL RST-COMPUTATION FOR CLOSURES ************************************) @@ -22,11 +22,11 @@ include "basic_2/rt_computation/fpbs_fqus.ma". lemma tdeq_fpbs_trans: ∀h,o,T1,T. T1 ≛[h, o] T → ∀G1,G2,L1,L2,T2. ⦃G1, L1, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄. -/3 width=5 by ffdeq_fpbs_trans, tdeq_ffdeq/ qed-. +/3 width=5 by fdeq_fpbs_trans, tdeq_fdeq/ qed-. lemma fpbs_tdeq_trans: ∀h,o,G1,G2,L1,L2,T1,T. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T⦄ → ∀T2. T ≛[h, o] T2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄. -/3 width=5 by fpbs_ffdeq_trans, tdeq_ffdeq/ qed-. +/3 width=5 by fpbs_fdeq_trans, tdeq_fdeq/ qed-. (* Properties with plus-iterated structural successor for closures **********) diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_lpxs.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_lpxs.ma index 26df6e480..3a67b6f49 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_lpxs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_lpxs.ma @@ -12,11 +12,11 @@ (* *) (**************************************************************************) -include "basic_2/static/ffdeq_fqus.ma". -include "basic_2/static/ffdeq_ffdeq.ma". +include "basic_2/static/fdeq_fqus.ma". +include "basic_2/static/fdeq_fdeq.ma". include "basic_2/rt_computation/cpxs_fqus.ma". -include "basic_2/rt_computation/cpxs_ffdeq.ma". -include "basic_2/rt_computation/lpxs_ffdeq.ma". +include "basic_2/rt_computation/cpxs_fdeq.ma". +include "basic_2/rt_computation/lpxs_fdeq.ma". include "basic_2/rt_computation/fpbs_cpxs.ma". (* PARALLEL RST-COMPUTATION FOR CLOSURES ************************************) @@ -41,9 +41,9 @@ lemma lpxs_fpbs_trans: ∀h,o,G1,G2,L,L2,T1,T2. ⦃G1, L, T1⦄ ≥[h, o] ⦃G2, qed-. (* Basic_2A1: uses: lpxs_lleq_fpbs *) -lemma lpxs_ffdeq_fpbs: ∀h,o,G1,L1,L,T1. ⦃G1, L1⦄ ⊢ ⬈*[h] L → - ∀G2,L2,T2. ⦃G1, L, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄. -/3 width=3 by lpxs_fpbs_trans, ffdeq_fpbs/ qed. +lemma lpxs_fdeq_fpbs: ∀h,o,G1,L1,L,T1. ⦃G1, L1⦄ ⊢ ⬈*[h] L → + ∀G2,L2,T2. ⦃G1, L, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄. +/3 width=3 by lpxs_fpbs_trans, fdeq_fpbs/ qed. lemma fpbs_lpx_trans: ∀h,o,G1,G2,L1,L,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L, T2⦄ → ∀L2. ⦃G2, L⦄ ⊢ ⬈[h] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄. @@ -79,7 +79,7 @@ lemma fpbs_intro_star: ∀h,o,G1,L1,T1,T. ⦃G1, L1⦄ ⊢ T1 ⬈*[h] T → ∀G,L,T0. ⦃G1, L1, T⦄ ⊐* ⦃G, L, T0⦄ → ∀L0. ⦃G, L⦄ ⊢ ⬈*[h] L0 → ∀G2,L2,T2. ⦃G, L0, T0⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ . -/3 width=5 by cpxs_fqus_lpxs_fpbs, fpbs_strap1, fpbq_ffdeq/ qed. +/3 width=5 by cpxs_fqus_lpxs_fpbs, fpbs_strap1, fpbq_fdeq/ qed. (* Advanced inversion lemmas *************************************************) @@ -100,10 +100,10 @@ lemma fpbs_inv_star: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, elim (lpx_fqus_trans … H34 … HL10) -L0 /3 width=9 by lpxs_step_sn, cpxs_trans, ex4_5_intro/ | #G0 #L0 #T0 #H10 #_ * #G3 #L3 #L4 #T3 #T4 #HT03 #H34 #HL34 #H42 - elim (ffdeq_cpxs_trans … H10 … HT03) -T0 #T0 #HT10 #H03 - elim (ffdeq_fqus_trans … H03 … H34) -G0 -L0 -T3 #G0 #L0 #T3 #H03 #H34 - elim (ffdeq_lpxs_trans … H34 … HL34) -L3 #L3 #HL03 #H34 - /3 width=13 by ffdeq_trans, ex4_5_intro/ + elim (fdeq_cpxs_trans … H10 … HT03) -T0 #T0 #HT10 #H03 + elim (fdeq_fqus_trans … H03 … H34) -G0 -L0 -T3 #G0 #L0 #T3 #H03 #H34 + elim (fdeq_lpxs_trans … H34 … HL34) -L3 #L3 #HL03 #H34 + /3 width=13 by fdeq_trans, ex4_5_intro/ ] ] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_csx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_csx.ma index ffcbc3d80..2746e5457 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_csx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_csx.ma @@ -47,7 +47,7 @@ generalize in match IHu; -IHu generalize in match H10; -H10 elim (lpx_fqup_trans … H03 … HL02) -L2 #L4 #T4 #HT04 #H43 #HL43 elim (tdeq_dec h o T0 T4) [ -IHc -HT04 #HT04 | -IHu #HnT04 ] [ elim (tdeq_fqup_trans … H43 … HT04) -T4 #L2 #T4 #H04 #HT43 #HL24 - /4 width=7 by fsb_fpbs_trans, tdeq_lfdeq_lpx_fpbs, fpbs_fqup_trans/ + /4 width=7 by fsb_fpbs_trans, tdeq_rdeq_lpx_fpbs, fpbs_fqup_trans/ | elim (cpxs_tdneq_fwd_step_sn … HT04 HnT04) -HT04 -HnT04 #T2 #T5 #HT02 #HnT02 #HT25 #HT54 elim (fpbs_cpx_tdneq_trans … H10 … HT02 HnT02) -T0 #T0 #HT10 #HnT10 #H02 /3 width=14 by fpbs_cpxs_tdeq_fqup_lpx_trans/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_fdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_fdeq.ma new file mode 100644 index 000000000..271a0e476 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_fdeq.ma @@ -0,0 +1,29 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/rt_transition/fpb_fdeq.ma". +include "basic_2/rt_computation/fsb.ma". + +(* STRONGLY NORMALIZING CLOSURES FOR PARALLEL RST-TRANSITION ****************) + +(* Properties with degree-based equivalence for closures ********************) + +lemma fsb_fdeq_trans: ∀h,o,G1,L1,T1. ≥[h, o] 𝐒⦃G1, L1, T1⦄ → + ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ≥[h, o] 𝐒⦃G2, L2, T2⦄. +#h #o #G1 #L1 #T1 #H @(fsb_ind_alt … H) -G1 -L1 -T1 +#G1 #L1 #T1 #_ #IH #G2 #L2 #T2 #H12 +@fsb_intro #G #L #T #H2 +elim (fdeq_fpb_trans … H12 … H2) -G2 -L2 -T2 +/2 width=5 by/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_ffdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_ffdeq.ma deleted file mode 100644 index fa0242a51..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_ffdeq.ma +++ /dev/null @@ -1,29 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/rt_transition/fpb_ffdeq.ma". -include "basic_2/rt_computation/fsb.ma". - -(* STRONGLY NORMALIZING CLOSURES FOR PARALLEL RST-TRANSITION ****************) - -(* Properties with degree-based equivalence for closures ********************) - -lemma fsb_ffdeq_trans: ∀h,o,G1,L1,T1. ≥[h, o] 𝐒⦃G1, L1, T1⦄ → - ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ≥[h, o] 𝐒⦃G2, L2, T2⦄. -#h #o #G1 #L1 #T1 #H @(fsb_ind_alt … H) -G1 -L1 -T1 -#G1 #L1 #T1 #_ #IH #G2 #L2 #T2 #H12 -@fsb_intro #G #L #T #H2 -elim (ffdeq_fpb_trans … H12 … H2) -G2 -L2 -T2 -/2 width=5 by/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_fpbg.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_fpbg.ma index b2093e068..07a9aad5e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_fpbg.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_fpbg.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "basic_2/rt_computation/fpbg_fpbs.ma". -include "basic_2/rt_computation/fsb_ffdeq.ma". +include "basic_2/rt_computation/fsb_fdeq.ma". (* STRONGLY NORMALIZING CLOSURES FOR PARALLEL RST-TRANSITION ****************) @@ -24,7 +24,7 @@ lemma fsb_fpbs_trans: ∀h,o,G1,L1,T1. ≥[h, o] 𝐒⦃G1, L1, T1⦄ → #h #o #G1 #L1 #T1 #H @(fsb_ind_alt … H) -G1 -L1 -T1 #G1 #L1 #T1 #H1 #IH #G2 #L2 #T2 #H12 elim (fpbs_inv_fpbg … H12) -H12 -[ -IH /2 width=5 by fsb_ffdeq_trans/ +[ -IH /2 width=5 by fsb_fdeq_trans/ | -H1 * /2 width=5 by/ ] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lpxs_fdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lpxs_fdeq.ma new file mode 100644 index 000000000..3a280f96f --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lpxs_fdeq.ma @@ -0,0 +1,29 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/fdeq.ma". +include "basic_2/rt_computation/lpxs_rdeq.ma". + +(* UNBOUND PARALLEL RT-COMPUTATION FOR FULL LOCAL ENVIRONMENTS **************) + +(* Properties with degree-based equivalence on closures *********************) + +lemma fdeq_lpxs_trans (h) (o): ∀G1,G2,L1,L0,T1,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L0, T2⦄ → + ∀L2. ⦃G2, L0⦄ ⊢⬈*[h] L2 → + ∃∃L. ⦃G1, L1⦄ ⊢⬈*[h] L & ⦃G1, L, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄. +#h #o #G1 #G2 #L1 #L0 #T1 #T2 #H1 #L2 #HL02 +elim (fdeq_inv_gen_dx … H1) -H1 #HG #HL10 #HT12 destruct +elim (rdeq_lpxs_trans … HL02 … HL10) -L0 #L0 #HL10 #HL02 +/3 width=3 by fdeq_intro_dx, ex2_intro/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lpxs_ffdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lpxs_ffdeq.ma deleted file mode 100644 index 530c5e15c..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lpxs_ffdeq.ma +++ /dev/null @@ -1,29 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/ffdeq.ma". -include "basic_2/rt_computation/lpxs_lfdeq.ma". - -(* UNBOUND PARALLEL RT-COMPUTATION FOR FULL LOCAL ENVIRONMENTS **************) - -(* Properties with degree-based equivalence on closures *********************) - -lemma ffdeq_lpxs_trans (h) (o): ∀G1,G2,L1,L0,T1,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L0, T2⦄ → - ∀L2. ⦃G2, L0⦄ ⊢⬈*[h] L2 → - ∃∃L. ⦃G1, L1⦄ ⊢⬈*[h] L & ⦃G1, L, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄. -#h #o #G1 #G2 #L1 #L0 #T1 #T2 #H1 #L2 #HL02 -elim (ffdeq_inv_gen_dx … H1) -H1 #HG #HL10 #HT12 destruct -elim (lfdeq_lpxs_trans … HL02 … HL10) -L0 #L0 #HL10 #HL02 -/3 width=3 by ffdeq_intro_dx, ex2_intro/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lpxs_lfdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lpxs_lfdeq.ma deleted file mode 100644 index 640c52779..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lpxs_lfdeq.ma +++ /dev/null @@ -1,50 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/rt_transition/lpx_lfdeq.ma". -include "basic_2/rt_computation/lpxs_lpx.ma". - -(* UNBOUND PARALLEL RT-COMPUTATION FOR FULL LOCAL ENVIRONMENTS **************) - -(* Properties with degree-based equivalence on referred entries *************) - -(* Basic_2A1: uses: lleq_lpxs_trans *) -lemma lfdeq_lpxs_trans (h) (o) (G) (T:term): ∀L2,K2. ⦃G, L2⦄ ⊢ ⬈*[h] K2 → - ∀L1. L1 ≛[h, o, T] L2 → - ∃∃K1. ⦃G, L1⦄ ⊢ ⬈*[h] K1 & K1 ≛[h, o, T] K2. -#h #o #G #T #L2 #K2 #H @(lpxs_ind_sn … H) -L2 /2 width=3 by ex2_intro/ -#L #L2 #HL2 #_ #IH #L1 #HT -elim (lfdeq_lpx_trans … HL2 … HT) -L #L #HL1 #HT -elim (IH … HT) -L2 #K #HLK #HT -/3 width=3 by lpxs_step_sn, ex2_intro/ -qed-. - -(* Basic_2A1: uses: lpxs_nlleq_inv_step_sn *) -lemma lpxs_lfdneq_inv_step_sn (h) (o) (G) (T:term): - ∀L1,L2. ⦃G, L1⦄ ⊢ ⬈*[h] L2 → (L1 ≛[h, o, T] L2 → ⊥) → - ∃∃L,L0. ⦃G, L1⦄ ⊢ ⬈[h] L & L1 ≛[h, o, T] L → ⊥ & - ⦃G, L⦄ ⊢ ⬈*[h] L0 & L0 ≛[h, o, T] L2. -#h #o #G #T #L1 #L2 #H @(lpxs_ind_sn … H) -L1 -[ #H elim H -H // -| #L1 #L #H1 #H2 #IH2 #H12 elim (lfdeq_dec h o L1 L T) #H - [ -H1 -H2 elim IH2 -IH2 /3 width=3 by lfdeq_trans/ -H12 - #L0 #L3 #H1 #H2 #H3 #H4 lapply (lfdeq_lfdneq_trans … H … H2) -H2 - #H2 elim (lfdeq_lpx_trans … H1 … H) -L - #L #H1 #H lapply (lfdneq_lfdeq_div … H … H2) -H2 - #H2 elim (lfdeq_lpxs_trans … H3 … H) -L0 - /3 width=8 by lfdeq_trans, ex4_2_intro/ - | -H12 -IH2 /3 width=6 by ex4_2_intro/ - ] -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lpxs_rdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lpxs_rdeq.ma new file mode 100644 index 000000000..3526c22c2 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lpxs_rdeq.ma @@ -0,0 +1,50 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/rt_transition/lpx_rdeq.ma". +include "basic_2/rt_computation/lpxs_lpx.ma". + +(* UNBOUND PARALLEL RT-COMPUTATION FOR FULL LOCAL ENVIRONMENTS **************) + +(* Properties with degree-based equivalence on referred entries *************) + +(* Basic_2A1: uses: lleq_lpxs_trans *) +lemma rdeq_lpxs_trans (h) (o) (G) (T:term): ∀L2,K2. ⦃G, L2⦄ ⊢ ⬈*[h] K2 → + ∀L1. L1 ≛[h, o, T] L2 → + ∃∃K1. ⦃G, L1⦄ ⊢ ⬈*[h] K1 & K1 ≛[h, o, T] K2. +#h #o #G #T #L2 #K2 #H @(lpxs_ind_sn … H) -L2 /2 width=3 by ex2_intro/ +#L #L2 #HL2 #_ #IH #L1 #HT +elim (rdeq_lpx_trans … HL2 … HT) -L #L #HL1 #HT +elim (IH … HT) -L2 #K #HLK #HT +/3 width=3 by lpxs_step_sn, ex2_intro/ +qed-. + +(* Basic_2A1: uses: lpxs_nlleq_inv_step_sn *) +lemma lpxs_rdneq_inv_step_sn (h) (o) (G) (T:term): + ∀L1,L2. ⦃G, L1⦄ ⊢ ⬈*[h] L2 → (L1 ≛[h, o, T] L2 → ⊥) → + ∃∃L,L0. ⦃G, L1⦄ ⊢ ⬈[h] L & L1 ≛[h, o, T] L → ⊥ & + ⦃G, L⦄ ⊢ ⬈*[h] L0 & L0 ≛[h, o, T] L2. +#h #o #G #T #L1 #L2 #H @(lpxs_ind_sn … H) -L1 +[ #H elim H -H // +| #L1 #L #H1 #H2 #IH2 #H12 elim (rdeq_dec h o L1 L T) #H + [ -H1 -H2 elim IH2 -IH2 /3 width=3 by rdeq_trans/ -H12 + #L0 #L3 #H1 #H2 #H3 #H4 lapply (rdeq_rdneq_trans … H … H2) -H2 + #H2 elim (rdeq_lpx_trans … H1 … H) -L + #L #H1 #H lapply (rdneq_rdeq_div … H … H2) -H2 + #H2 elim (rdeq_lpxs_trans … H3 … H) -L0 + /3 width=8 by rdeq_trans, ex4_2_intro/ + | -H12 -IH2 /3 width=6 by ex4_2_intro/ + ] +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lsubsx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lsubsx.ma index 66104ff6a..0f4601234 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lsubsx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lsubsx.ma @@ -17,7 +17,7 @@ include "basic_2/rt_computation/rdsx.ma". (* CLEAR OF STRONGLY NORMALIZING ENTRIES FOR UNBOUND RT-TRANSITION **********) -(* Note: this should be an instance of a more general lexs *) +(* Note: this should be an instance of a more general sex *) (* Basic_2A1: uses: lcosx *) inductive lsubsx (h) (o) (G): rtmap → relation lenv ≝ | lsubsx_atom: ∀f. lsubsx h o G f (⋆) (⋆) diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx.ma index 4503f5e3e..3ec296cfa 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx.ma @@ -13,13 +13,13 @@ (**************************************************************************) include "basic_2/notation/relations/predtysnstrong_5.ma". -include "basic_2/static/lfdeq.ma". +include "basic_2/static/rdeq.ma". include "basic_2/rt_transition/lpx.ma". (* STRONGLY NORMALIZING REFERRED LOCAL ENV.S FOR UNBOUND RT-TRANSITION ******) definition rdsx (h) (o) (G) (T): predicate lenv ≝ - SN … (lpx h G) (lfdeq h o T). + SN … (lpx h G) (rdeq h o T). interpretation "strong normalization for unbound context-sensitive parallel rt-transition on referred entries (local environment)" @@ -57,7 +57,7 @@ lemma rdsx_fwd_pair_sn (h) (o) (G): #h #o #G #I #L #V #T #H @(rdsx_ind … H) -L #L1 #_ #IHL1 @rdsx_intro #L2 #HL12 #HnL12 -/4 width=3 by lfdeq_fwd_pair_sn/ +/4 width=3 by rdeq_fwd_pair_sn/ qed-. (* Basic_2A1: uses: lsx_fwd_flat_dx *) @@ -67,14 +67,14 @@ lemma rdsx_fwd_flat_dx (h) (o) (G): #h #o #G #I #L #V #T #H @(rdsx_ind … H) -L #L1 #_ #IHL1 @rdsx_intro #L2 #HL12 #HnL12 -/4 width=3 by lfdeq_fwd_flat_dx/ +/4 width=3 by rdeq_fwd_flat_dx/ qed-. fact rdsx_fwd_pair_aux (h) (o) (G): ∀L. G ⊢ ⬈*[h, o, #0] 𝐒⦃L⦄ → ∀I,K,V. L = K.ⓑ{I}V → G ⊢ ⬈*[h, o, V] 𝐒⦃K⦄. #h #o #G #L #H @(rdsx_ind … H) -L #L1 #_ #IH #I #K1 #V #H destruct -/5 width=5 by lpx_pair, rdsx_intro, lfdeq_fwd_zero_pair/ +/5 width=5 by lpx_pair, rdsx_intro, rdeq_fwd_zero_pair/ qed-. lemma rdsx_fwd_pair (h) (o) (G): diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_csx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_csx.ma index 390ebe3ce..355ecb004 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_csx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_csx.ma @@ -31,7 +31,7 @@ lemma rdsx_pair_lpxs (h) (o) (G): @rdsx_intro #Y #HY #HnY elim (lpx_inv_pair_sn … HY) -HY #K2 #V2 #HK02 #HV02 #H destruct elim (tdeq_dec h o V0 V2) #HnV02 destruct [ -IHV0 -HV02 -HK0 | -IHK0 -HnY ] -[ /5 width=5 by rdsx_lfdeq_trans, lpxs_step_dx, lfdeq_pair/ +[ /5 width=5 by rdsx_rdeq_trans, lpxs_step_dx, rdeq_pair/ | @(IHV0 … HnV02) -IHV0 -HnV02 [ /2 width=3 by lpxs_cpx_trans/ | /3 width=3 by rdsx_lpx_trans, rdsx_cpx_trans/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_drops.ma index db1bd5144..6f0cc4da0 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_drops.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/static/lfdeq_drops.ma". +include "basic_2/static/rdeq_drops.ma". include "basic_2/rt_transition/lpx_drops.ma". include "basic_2/rt_computation/rdsx_length.ma". include "basic_2/rt_computation/rdsx_fqup.ma". @@ -27,7 +27,7 @@ lemma rdsx_lifts (h) (o) (G): d_liftable1_isuni … (λL,T. G ⊢ ⬈*[h, o, T] #h #o #G #K #T #H @(rdsx_ind … H) -K #K1 #_ #IH #b #f #L1 #HLK1 #Hf #U #HTU @rdsx_intro #L2 #HL12 #HnL12 elim (lpx_drops_conf … HLK1 … HL12) -/5 width=9 by lfdeq_lifts_bi, lpx_fwd_length/ +/5 width=9 by rdeq_lifts_bi, lpx_fwd_length/ qed-. (* Inversion lemmas on relocation *******************************************) @@ -37,7 +37,7 @@ lemma rdsx_inv_lifts (h) (o) (G): d_deliftable1_isuni … (λL,T. G ⊢ ⬈*[h, #h #o #G #L #U #H @(rdsx_ind … H) -L #L1 #_ #IH #b #f #K1 #HLK1 #Hf #T #HTU @rdsx_intro #K2 #HK12 #HnK12 elim (drops_lpx_trans … HLK1 … HK12) -HK12 -/4 width=10 by lfdeq_inv_lifts_bi/ +/4 width=10 by rdeq_inv_lifts_bi/ qed-. (* Advanced properties ******************************************************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_fqup.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_fqup.ma index 2748790e7..6216111e0 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_fqup.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_fqup.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/static/lfdeq_fqup.ma". +include "basic_2/static/rdeq_fqup.ma". include "basic_2/rt_computation/rdsx.ma". (* STRONGLY NORMALIZING REFERRED LOCAL ENV.S FOR UNBOUND RT-TRANSITION ******) @@ -39,7 +39,7 @@ lemma rdsx_fwd_bind_dx (h) (o) (G): @(rdsx_ind … H) -L #L1 #_ #IH @rdsx_intro #Y #H #HT elim (lpx_inv_unit_sn … H) -H #L2 #HL12 #H destruct -/4 width=4 by lfdeq_fwd_bind_dx_void/ +/4 width=4 by rdeq_fwd_bind_dx_void/ qed-. (* Advanced inversion lemmas ************************************************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_length.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_length.ma index 7063209b5..e6ccab06a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_length.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_length.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/static/lfdeq_length.ma". +include "basic_2/static/rdeq_length.ma". include "basic_2/rt_transition/lpx_length.ma". include "basic_2/rt_computation/rdsx.ma". @@ -23,18 +23,18 @@ include "basic_2/rt_computation/rdsx.ma". (* Basic_2A1: uses: lsx_sort *) lemma rdsx_sort (h) (o) (G): ∀L,s. G ⊢ ⬈*[h, o, ⋆s] 𝐒⦃L⦄. #h #o #G #L1 #s @rdsx_intro #L2 #H #Hs -elim Hs -Hs /3 width=3 by lpx_fwd_length, lfdeq_sort_length/ +elim Hs -Hs /3 width=3 by lpx_fwd_length, rdeq_sort_length/ qed. (* Basic_2A1: uses: lsx_gref *) lemma rdsx_gref (h) (o) (G): ∀L,l. G ⊢ ⬈*[h, o, §l] 𝐒⦃L⦄. #h #o #G #L1 #s @rdsx_intro #L2 #H #Hs -elim Hs -Hs /3 width=3 by lpx_fwd_length, lfdeq_gref_length/ +elim Hs -Hs /3 width=3 by lpx_fwd_length, rdeq_gref_length/ qed. lemma rdsx_unit (h) (o) (G): ∀I,L. G ⊢ ⬈*[h, o, #0] 𝐒⦃L.ⓤ{I}⦄. #h #o #G #I #L1 @rdsx_intro #Y #HY #HnY elim HnY -HnY elim (lpx_inv_unit_sn … HY) -HY #L2 #HL12 #H destruct -/3 width=3 by lpx_fwd_length, lfdeq_unit_length/ +/3 width=3 by lpx_fwd_length, rdeq_unit_length/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_lpxs.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_lpxs.ma index c9f49cfe1..06649215c 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_lpxs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_lpxs.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/rt_computation/lpxs_lfdeq.ma". +include "basic_2/rt_computation/lpxs_rdeq.ma". include "basic_2/rt_computation/lpxs_lpxs.ma". include "basic_2/rt_computation/rdsx_rdsx.ma". @@ -35,28 +35,28 @@ qed-. (* Eliminators with unbound rt-computation for full local environments ******) -lemma rdsx_ind_lpxs_lfdeq (h) (o) (G): - ∀T. ∀Q:predicate lenv. - (∀L1. G ⊢ ⬈*[h, o, T] 𝐒⦃L1⦄ → - (∀L2. ⦃G, L1⦄ ⊢ ⬈*[h] L2 → (L1 ≛[h, o, T] L2 → ⊥) → Q L2) → - Q L1 - ) → - ∀L1. G ⊢ ⬈*[h, o, T] 𝐒⦃L1⦄ → - ∀L0. ⦃G, L1⦄ ⊢ ⬈*[h] L0 → ∀L2. L0 ≛[h, o, T] L2 → Q L2. +lemma rdsx_ind_lpxs_rdeq (h) (o) (G): + ∀T. ∀Q:predicate lenv. + (∀L1. G ⊢ ⬈*[h, o, T] 𝐒⦃L1⦄ → + (∀L2. ⦃G, L1⦄ ⊢ ⬈*[h] L2 → (L1 ≛[h, o, T] L2 → ⊥) → Q L2) → + Q L1 + ) → + ∀L1. G ⊢ ⬈*[h, o, T] 𝐒⦃L1⦄ → + ∀L0. ⦃G, L1⦄ ⊢ ⬈*[h] L0 → ∀L2. L0 ≛[h, o, T] L2 → Q L2. #h #o #G #T #Q #IH #L1 #H @(rdsx_ind … H) -L1 #L1 #HL1 #IH1 #L0 #HL10 #L2 #HL02 -@IH -IH /3 width=3 by rdsx_lpxs_trans, rdsx_lfdeq_trans/ -HL1 #K2 #HLK2 #HnLK2 -lapply (lfdeq_lfdneq_trans … HL02 … HnLK2) -HnLK2 #H -elim (lfdeq_lpxs_trans … HLK2 … HL02) -L2 #K0 #HLK0 #HK02 -lapply (lfdneq_lfdeq_canc_dx … H … HK02) -H #HnLK0 -elim (lfdeq_dec h o L1 L0 T) #H -[ lapply (lfdeq_lfdneq_trans … H … HnLK0) -H -HnLK0 #Hn10 +@IH -IH /3 width=3 by rdsx_lpxs_trans, rdsx_rdeq_trans/ -HL1 #K2 #HLK2 #HnLK2 +lapply (rdeq_rdneq_trans … HL02 … HnLK2) -HnLK2 #H +elim (rdeq_lpxs_trans … HLK2 … HL02) -L2 #K0 #HLK0 #HK02 +lapply (rdneq_rdeq_canc_dx … H … HK02) -H #HnLK0 +elim (rdeq_dec h o L1 L0 T) #H +[ lapply (rdeq_rdneq_trans … H … HnLK0) -H -HnLK0 #Hn10 lapply (lpxs_trans … HL10 … HLK0) -L0 #H10 - elim (lpxs_lfdneq_inv_step_sn … H10 … Hn10) -H10 -Hn10 - /3 width=8 by lfdeq_trans/ -| elim (lpxs_lfdneq_inv_step_sn … HL10 … H) -HL10 -H #L #K #HL1 #HnL1 #HLK #HKL0 - elim (lfdeq_lpxs_trans … HLK0 … HKL0) -L0 - /3 width=8 by lpxs_trans, lfdeq_trans/ + elim (lpxs_rdneq_inv_step_sn … H10 … Hn10) -H10 -Hn10 + /3 width=8 by rdeq_trans/ +| elim (lpxs_rdneq_inv_step_sn … HL10 … H) -HL10 -H #L #K #HL1 #HnL1 #HLK #HKL0 + elim (rdeq_lpxs_trans … HLK0 … HKL0) -L0 + /3 width=8 by lpxs_trans, rdeq_trans/ ] qed-. @@ -69,7 +69,7 @@ lemma rdsx_ind_lpxs (h) (o) (G): ) → ∀L. G ⊢ ⬈*[h, o, T] 𝐒⦃L⦄ → Q L. #h #o #G #T #Q #IH #L #HL -@(rdsx_ind_lpxs_lfdeq … IH … HL) -IH -HL // (**) (* full auto fails *) +@(rdsx_ind_lpxs_rdeq … IH … HL) -IH -HL // (**) (* full auto fails *) qed-. (* Advanced properties ******************************************************) @@ -84,10 +84,10 @@ fact rdsx_bind_lpxs_aux (h) (o) (G): #Y #HY #IHY #L2 #H #HL12 destruct @rdsx_intro_lpxs #L0 #HL20 lapply (lpxs_trans … HL12 … HL20) #HL10 #H -elim (lfdneq_inv_bind … H) -H [ -IHY | -HY -IHL1 -HL12 ] -[ #HnV elim (lfdeq_dec h o L1 L2 V) +elim (rdneq_inv_bind … H) -H [ -IHY | -HY -IHL1 -HL12 ] +[ #HnV elim (rdeq_dec h o L1 L2 V) [ #HV @(IHL1 … HL10) -IHL1 -HL12 -HL10 - /3 width=4 by rdsx_lpxs_trans, lpxs_bind_refl_dx, lfdeq_canc_sn/ (**) (* full auto too slow *) + /3 width=4 by rdsx_lpxs_trans, lpxs_bind_refl_dx, rdeq_canc_sn/ (**) (* full auto too slow *) | -HnV -HL10 /4 width=4 by rdsx_lpxs_trans, lpxs_bind_refl_dx/ ] | /3 width=4 by lpxs_bind_refl_dx/ @@ -110,10 +110,10 @@ lemma rdsx_flat_lpxs (h) (o) (G): #L1 #HL1 #IHL1 #L2 #T #H @(rdsx_ind_lpxs … H) -L2 #L2 #HL2 #IHL2 #HL12 @rdsx_intro_lpxs #L0 #HL20 lapply (lpxs_trans … HL12 … HL20) -#HL10 #H elim (lfdneq_inv_flat … H) -H [ -HL1 -IHL2 | -HL2 -IHL1 ] -[ #HnV elim (lfdeq_dec h o L1 L2 V) +#HL10 #H elim (rdneq_inv_flat … H) -H [ -HL1 -IHL2 | -HL2 -IHL1 ] +[ #HnV elim (rdeq_dec h o L1 L2 V) [ #HV @(IHL1 … HL10) -IHL1 -HL12 -HL10 - /3 width=5 by rdsx_lpxs_trans, lfdeq_canc_sn/ (**) (* full auto too slow: 47s *) + /3 width=5 by rdsx_lpxs_trans, rdeq_canc_sn/ (**) (* full auto too slow: 47s *) | -HnV -HL10 /3 width=4 by rdsx_lpxs_trans/ ] | /3 width=3 by/ @@ -136,10 +136,10 @@ fact rdsx_bind_lpxs_void_aux (h) (o) (G): #Y #HY #IHY #L2 #H #HL12 destruct @rdsx_intro_lpxs #L0 #HL20 lapply (lpxs_trans … HL12 … HL20) #HL10 #H -elim (lfdneq_inv_bind_void … H) -H [ -IHY | -HY -IHL1 -HL12 ] -[ #HnV elim (lfdeq_dec h o L1 L2 V) +elim (rdneq_inv_bind_void … H) -H [ -IHY | -HY -IHL1 -HL12 ] +[ #HnV elim (rdeq_dec h o L1 L2 V) [ #HV @(IHL1 … HL10) -IHL1 -HL12 -HL10 - /3 width=6 by rdsx_lpxs_trans, lpxs_bind_refl_dx, lfdeq_canc_sn/ (**) (* full auto too slow *) + /3 width=6 by rdsx_lpxs_trans, lpxs_bind_refl_dx, rdeq_canc_sn/ (**) (* full auto too slow *) | -HnV -HL10 /4 width=4 by rdsx_lpxs_trans, lpxs_bind_refl_dx/ ] | /3 width=4 by lpxs_bind_refl_dx/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_rdsx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_rdsx.ma index efdf72af6..2305aac19 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_rdsx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/rdsx_rdsx.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/rt_transition/lpx_lfdeq.ma". +include "basic_2/rt_transition/lpx_rdeq.ma". include "basic_2/rt_computation/rdsx.ma". (* STRONGLY NORMALIZING REFERRED LOCAL ENV.S FOR UNBOUND RT-TRANSITION ******) @@ -20,13 +20,13 @@ include "basic_2/rt_computation/rdsx.ma". (* Advanced properties ******************************************************) (* Basic_2A1: uses: lsx_lleq_trans *) -lemma rdsx_lfdeq_trans (h) (o) (G): - ∀L1,T. G ⊢ ⬈*[h, o, T] 𝐒⦃L1⦄ → - ∀L2. L1 ≛[h, o, T] L2 → G ⊢ ⬈*[h, o, T] 𝐒⦃L2⦄. +lemma rdsx_rdeq_trans (h) (o) (G): + ∀L1,T. G ⊢ ⬈*[h, o, T] 𝐒⦃L1⦄ → + ∀L2. L1 ≛[h, o, T] L2 → G ⊢ ⬈*[h, o, T] 𝐒⦃L2⦄. #h #o #G #L1 #T #H @(rdsx_ind … H) -L1 #L1 #_ #IHL1 #L2 #HL12 @rdsx_intro -#L #HL2 #HnL2 elim (lfdeq_lpx_trans … HL2 … HL12) -HL2 -/4 width=5 by lfdeq_repl/ +#L #HL2 #HnL2 elim (rdeq_lpx_trans … HL2 … HL12) -HL2 +/4 width=5 by rdeq_repl/ qed-. (* Basic_2A1: uses: lsx_lpx_trans *) @@ -34,5 +34,5 @@ lemma rdsx_lpx_trans (h) (o) (G): ∀L1,T. G ⊢ ⬈*[h, o, T] 𝐒⦃L1⦄ → ∀L2. ⦃G, L1⦄ ⊢ ⬈[h] L2 → G ⊢ ⬈*[h, o, T] 𝐒⦃L2⦄. #h #o #G #L1 #T #H @(rdsx_ind … H) -L1 #L1 #HL1 #IHL1 #L2 #HL12 -elim (lfdeq_dec h o L1 L2 T) /3 width=4 by rdsx_lfdeq_trans/ +elim (rdeq_dec h o L1 L2 T) /3 width=4 by rdsx_rdeq_trans/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cnx_cnx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cnx_cnx.ma index 4fd229e59..1cd716e8e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cnx_cnx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cnx_cnx.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/rt_transition/lfpx_lfdeq.ma". +include "basic_2/rt_transition/rpx_rdeq.ma". include "basic_2/rt_transition/cnx.ma". (* NORMAL TERMS FOR UNBOUND CONTEXT-SENSITIVE PARALLEL RT-TRANSITION ********) diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm_fsle.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm_fsle.ma index c9d76fc51..0c51efd06 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm_fsle.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm_fsle.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/rt_transition/lfpx_fsle.ma". +include "basic_2/rt_transition/rpx_fsle.ma". include "basic_2/rt_transition/cpm_cpx.ma". (* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************) @@ -22,11 +22,11 @@ include "basic_2/rt_transition/cpm_cpx.ma". lemma cpm_fsge_comp: ∀n,h,G. R_fsge_compatible (λL. cpm h G L n). /3 width=6 by cpm_fwd_cpx, cpx_fsge_comp/ qed-. -lemma lfpr_fsge_comp: ∀h,G. lfxs_fsge_compatible (λL. cpm h G L 0). -/4 width=5 by cpm_fwd_cpx, lfpx_fsge_comp, lfxs_co/ qed-. +lemma rpr_fsge_comp: ∀h,G. rex_fsge_compatible (λL. cpm h G L 0). +/4 width=5 by cpm_fwd_cpx, rpx_fsge_comp, rex_co/ qed-. (* Properties with generic extension on referred entries ********************) (* Basic_2A1: was just: cpr_llpx_sn_conf *) -lemma cpm_lfxs_conf: ∀R,n,h,G. s_r_confluent1 … (λL. cpm h G L n) (lfxs R). -/3 width=5 by cpm_fwd_cpx, cpx_lfxs_conf/ qed-. +lemma cpm_rex_conf: ∀R,n,h,G. s_r_confluent1 … (λL. cpm h G L n) (rex R). +/3 width=5 by cpm_fwd_cpx, cpx_rex_conf/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_fdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_fdeq.ma new file mode 100644 index 000000000..3a86118e1 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_fdeq.ma @@ -0,0 +1,32 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/fdeq.ma". +include "basic_2/rt_transition/cpx_rdeq.ma". +include "basic_2/rt_transition/rpx_rdeq.ma". + +(* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS ***************) + +(* Properties with degree-based equivalence for closures ********************) + +lemma fdeq_cpx_trans: ∀h,o,G1,G2,L1,L2,T1,T. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T⦄ → + ∀T2. ⦃G2, L2⦄ ⊢ T ⬈[h] T2 → + ∃∃T0. ⦃G1, L1⦄ ⊢ T1 ⬈[h] T0 & ⦃G1, L1, T0⦄ ≛[h, o] ⦃G2, L2, T2⦄. +#h #o #G1 #G2 #L1 #L2 #T1 #T #H #T2 #HT2 +elim (fdeq_inv_gen_dx … H) -H #H #HL12 #HT1 destruct +elim (rdeq_cpx_trans … HL12 … HT2) #T0 #HT0 #HT02 +lapply (cpx_rdeq_conf_dx … HT2 … HL12) -HL12 #HL12 +elim (tdeq_cpx_trans … HT1 … HT0) -T #T #HT1 #HT0 +/4 width=5 by fdeq_intro_dx, tdeq_trans, ex2_intro/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_ffdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_ffdeq.ma deleted file mode 100644 index decfa484a..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_ffdeq.ma +++ /dev/null @@ -1,32 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/ffdeq.ma". -include "basic_2/rt_transition/cpx_lfdeq.ma". -include "basic_2/rt_transition/lfpx_lfdeq.ma". - -(* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS ***************) - -(* Properties with degree-based equivalence for closures ********************) - -lemma ffdeq_cpx_trans: ∀h,o,G1,G2,L1,L2,T1,T. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T⦄ → - ∀T2. ⦃G2, L2⦄ ⊢ T ⬈[h] T2 → - ∃∃T0. ⦃G1, L1⦄ ⊢ T1 ⬈[h] T0 & ⦃G1, L1, T0⦄ ≛[h, o] ⦃G2, L2, T2⦄. -#h #o #G1 #G2 #L1 #L2 #T1 #T #H #T2 #HT2 -elim (ffdeq_inv_gen_dx … H) -H #H #HL12 #HT1 destruct -elim (lfdeq_cpx_trans … HL12 … HT2) #T0 #HT0 #HT02 -lapply (cpx_lfdeq_conf_dx … HT2 … HL12) -HL12 #HL12 -elim (tdeq_cpx_trans … HT1 … HT0) -T #T #HT1 #HT0 -/4 width=5 by ffdeq_intro_dx, tdeq_trans, ex2_intro/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_lfdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_lfdeq.ma deleted file mode 100644 index 45a707603..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_lfdeq.ma +++ /dev/null @@ -1,29 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfdeq_lfdeq.ma". -include "basic_2/rt_transition/lfpx_fsle.ma". - -(* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS ***************) - -(* Properties with degree-based equivalence for local environments **********) - -(* Basic_2A1: was just: cpx_lleq_conf_sn *) -lemma cpx_lfdeq_conf_sn: ∀h,o,G. s_r_confluent1 … (cpx h G) (lfdeq h o). -/3 width=6 by cpx_lfxs_conf/ qed-. - -(* Basic_2A1: was just: cpx_lleq_conf_dx *) -lemma cpx_lfdeq_conf_dx: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ⬈[h] T2 → - ∀L1. L1 ≛[h, o, T1] L2 → L1 ≛[h, o, T2] L2. -/4 width=4 by cpx_lfdeq_conf_sn, lfdeq_sym/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_lfeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_lfeq.ma deleted file mode 100644 index 48ae7a970..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_lfeq.ma +++ /dev/null @@ -1,63 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfeq.ma". -include "basic_2/rt_transition/lfpx_fsle.ma". - -(* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS ***************) - -(* Properties with syntactic equivalence for lenvs on referred entries ******) - -(* Basic_2A1: was: lleq_cpx_trans *) -lemma lfeq_cpx_trans: ∀h,G. lfeq_transitive (cpx h G). -#h #G #L2 #T1 #T2 #H @(cpx_ind … H) -G -L2 -T1 -T2 /2 width=2 by cpx_ess/ -[ #I #G #K2 #V1 #V2 #W2 #_ #IH #HVW2 #L1 #H - elim (lfeq_inv_zero_pair_dx … H) -H #K1 #HK12 #H destruct - /3 width=3 by cpx_delta/ -| #I2 #G #K2 #T #U #i #_ #IH #HTU #L1 #H - elim (lfeq_inv_lref_bind_dx … H) -H #I1 #K1 #HK12 #H destruct - /3 width=3 by cpx_lref/ -| #p #I #G #L2 #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #H - elim (lfeq_inv_bind … H) -H /3 width=1 by cpx_bind/ -| #I #G #L2 #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #H - elim (lfeq_inv_flat … H) -H /3 width=1 by cpx_flat/ -| #G #L2 #V2 #T1 #T #T2 #_ #HT2 #IH #L1 #H - elim (lfeq_inv_bind … H) -H /3 width=3 by cpx_zeta/ -| #G #L2 #W1 #T1 #T2 #_ #IH #L1 #H - elim (lfeq_inv_flat … H) -H /3 width=1 by cpx_eps/ -| #G #L2 #W1 #W2 #T1 #_ #IH #L1 #H - elim (lfeq_inv_flat … H) -H /3 width=1 by cpx_ee/ -| #p #G #L2 #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #L1 #H - elim (lfeq_inv_flat … H) -H #HV1 #H - elim (lfeq_inv_bind … H) -H /3 width=1 by cpx_beta/ -| #p #G #L2 #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV1 #IHW12 #IHT12 #HV2 #L1 #H - elim (lfeq_inv_flat … H) -H #HV1 #H - elim (lfeq_inv_bind … H) -H /3 width=3 by cpx_theta/ -] -qed-. -(* -(* Basic_2A1: was: cpx_lleq_conf *) -lemma cpx_lfeq_conf: ∀h,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ⬈[h] T2 → - ∀L1. L2 ≘[T1] L1 → ⦃G, L1⦄ ⊢ T1 ⬈[h] T2. -/3 width=3 by lfeq_cpx_trans, lfeq_sym/ qed-. -*) -(* Basic_2A1: was: cpx_lleq_conf_sn *) -lemma cpx_lfeq_conf_sn: ∀h,G. s_r_confluent1 … (cpx h G) lfeq. -/2 width=5 by cpx_lfxs_conf/ qed-. -(* -(* Basic_2A1: was: cpx_lleq_conf_dx *) -lemma cpx_lfeq_conf_dx: ∀h,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ⬈[h] T2 → - ∀L1. L1 ≘[T1] L2 → L1 ≘[T2] L2. -/4 width=6 by cpx_lfeq_conf_sn, lfeq_sym/ qed-. -*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_rdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_rdeq.ma new file mode 100644 index 000000000..1f5e4a9bb --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_rdeq.ma @@ -0,0 +1,29 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rdeq_rdeq.ma". +include "basic_2/rt_transition/rpx_fsle.ma". + +(* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS ***************) + +(* Properties with degree-based equivalence for local environments **********) + +(* Basic_2A1: was just: cpx_lleq_conf_sn *) +lemma cpx_rdeq_conf_sn: ∀h,o,G. s_r_confluent1 … (cpx h G) (rdeq h o). +/3 width=6 by cpx_rex_conf/ qed-. + +(* Basic_2A1: was just: cpx_lleq_conf_dx *) +lemma cpx_rdeq_conf_dx: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ⬈[h] T2 → + ∀L1. L1 ≛[h, o, T1] L2 → L1 ≛[h, o, T2] L2. +/4 width=4 by cpx_rdeq_conf_sn, rdeq_sym/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_req.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_req.ma new file mode 100644 index 000000000..44dcfb53c --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_req.ma @@ -0,0 +1,63 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/req.ma". +include "basic_2/rt_transition/rpx_fsle.ma". + +(* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS ***************) + +(* Properties with syntactic equivalence for lenvs on referred entries ******) + +(* Basic_2A1: was: lleq_cpx_trans *) +lemma req_cpx_trans: ∀h,G. req_transitive (cpx h G). +#h #G #L2 #T1 #T2 #H @(cpx_ind … H) -G -L2 -T1 -T2 /2 width=2 by cpx_ess/ +[ #I #G #K2 #V1 #V2 #W2 #_ #IH #HVW2 #L1 #H + elim (req_inv_zero_pair_dx … H) -H #K1 #HK12 #H destruct + /3 width=3 by cpx_delta/ +| #I2 #G #K2 #T #U #i #_ #IH #HTU #L1 #H + elim (req_inv_lref_bind_dx … H) -H #I1 #K1 #HK12 #H destruct + /3 width=3 by cpx_lref/ +| #p #I #G #L2 #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #H + elim (req_inv_bind … H) -H /3 width=1 by cpx_bind/ +| #I #G #L2 #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #H + elim (req_inv_flat … H) -H /3 width=1 by cpx_flat/ +| #G #L2 #V2 #T1 #T #T2 #_ #HT2 #IH #L1 #H + elim (req_inv_bind … H) -H /3 width=3 by cpx_zeta/ +| #G #L2 #W1 #T1 #T2 #_ #IH #L1 #H + elim (req_inv_flat … H) -H /3 width=1 by cpx_eps/ +| #G #L2 #W1 #W2 #T1 #_ #IH #L1 #H + elim (req_inv_flat … H) -H /3 width=1 by cpx_ee/ +| #p #G #L2 #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #L1 #H + elim (req_inv_flat … H) -H #HV1 #H + elim (req_inv_bind … H) -H /3 width=1 by cpx_beta/ +| #p #G #L2 #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV1 #IHW12 #IHT12 #HV2 #L1 #H + elim (req_inv_flat … H) -H #HV1 #H + elim (req_inv_bind … H) -H /3 width=3 by cpx_theta/ +] +qed-. +(* +(* Basic_2A1: was: cpx_lleq_conf *) +lemma cpx_req_conf: ∀h,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ⬈[h] T2 → + ∀L1. L2 ≘[T1] L1 → ⦃G, L1⦄ ⊢ T1 ⬈[h] T2. +/3 width=3 by req_cpx_trans, req_sym/ qed-. +*) +(* Basic_2A1: was: cpx_lleq_conf_sn *) +lemma cpx_req_conf_sn: ∀h,G. s_r_confluent1 … (cpx h G) req. +/2 width=5 by cpx_rex_conf/ qed-. +(* +(* Basic_2A1: was: cpx_lleq_conf_dx *) +lemma cpx_req_conf_dx: ∀h,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ⬈[h] T2 → + ∀L1. L1 ≘[T1] L2 → L1 ≘[T2] L2. +/4 width=6 by cpx_req_conf_sn, req_sym/ qed-. +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpb.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpb.ma index 268e45283..6d163a685 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpb.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpb.ma @@ -14,7 +14,7 @@ include "basic_2/notation/relations/predsubtyproper_8.ma". include "basic_2/s_transition/fqu.ma". -include "basic_2/static/lfdeq.ma". +include "basic_2/static/rdeq.ma". include "basic_2/rt_transition/lpr_lpx.ma". (* PROPER PARALLEL RST-TRANSITION FOR CLOSURES ******************************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpb_fdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpb_fdeq.ma new file mode 100644 index 000000000..4aaf35624 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpb_fdeq.ma @@ -0,0 +1,46 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/s_transition/fqu_tdeq.ma". +include "basic_2/static/fdeq.ma". +include "basic_2/rt_transition/fpb_rdeq.ma". + +(* PROPER PARALLEL RST-TRANSITION FOR CLOSURES ******************************) + +(* Properties with degree-based equivalence for closures ********************) + +(* Basic_2A1: uses: fleq_fpb_trans *) +lemma fdeq_fpb_trans: ∀h,o,F1,F2,K1,K2,T1,T2. ⦃F1, K1, T1⦄ ≛[h, o] ⦃F2, K2, T2⦄ → + ∀G2,L2,U2. ⦃F2, K2, T2⦄ ≻[h, o] ⦃G2, L2, U2⦄ → + ∃∃G1,L1,U1. ⦃F1, K1, T1⦄ ≻[h, o] ⦃G1, L1, U1⦄ & ⦃G1, L1, U1⦄ ≛[h, o] ⦃G2, L2, U2⦄. +#h #o #F1 #F2 #K1 #K2 #T1 #T2 * -F2 -K2 -T2 +#K2 #T2 #HK12 #HT12 #G2 #L2 #U2 #H12 +elim (tdeq_fpb_trans … HT12 … H12) -T2 #K0 #T0 #H #HT0 #HK0 +elim (rdeq_fpb_trans … HK12 … H) -K2 #L0 #U0 #H #HUT0 #HLK0 +@(ex2_3_intro … H) -H (**) (* full auto too slow *) +/4 width=3 by fdeq_intro_dx, rdeq_trans, tdeq_rdeq_conf, tdeq_trans/ +qed-. + +(* Inversion lemmas with degree-based equivalence for closures **************) + +(* Basic_2A1: uses: fpb_inv_fleq *) +lemma fpb_inv_fdeq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → + ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⊥. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 +[ #G2 #L2 #T2 #H12 #H elim (fdeq_inv_gen_sn … H) -H + /3 width=11 by rdeq_fwd_length, fqu_inv_tdeq/ +| #T2 #_ #HnT #H elim (fdeq_inv_gen_sn … H) -H /2 width=1 by/ +| #L2 #_ #HnL #H elim (fdeq_inv_gen_sn … H) -H /2 width=1 by/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpb_ffdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpb_ffdeq.ma deleted file mode 100644 index 5052736d8..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpb_ffdeq.ma +++ /dev/null @@ -1,46 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/s_transition/fqu_tdeq.ma". -include "basic_2/static/ffdeq.ma". -include "basic_2/rt_transition/fpb_lfdeq.ma". - -(* PROPER PARALLEL RST-TRANSITION FOR CLOSURES ******************************) - -(* Properties with degree-based equivalence for closures ********************) - -(* Basic_2A1: uses: fleq_fpb_trans *) -lemma ffdeq_fpb_trans: ∀h,o,F1,F2,K1,K2,T1,T2. ⦃F1, K1, T1⦄ ≛[h, o] ⦃F2, K2, T2⦄ → - ∀G2,L2,U2. ⦃F2, K2, T2⦄ ≻[h, o] ⦃G2, L2, U2⦄ → - ∃∃G1,L1,U1. ⦃F1, K1, T1⦄ ≻[h, o] ⦃G1, L1, U1⦄ & ⦃G1, L1, U1⦄ ≛[h, o] ⦃G2, L2, U2⦄. -#h #o #F1 #F2 #K1 #K2 #T1 #T2 * -F2 -K2 -T2 -#K2 #T2 #HK12 #HT12 #G2 #L2 #U2 #H12 -elim (tdeq_fpb_trans … HT12 … H12) -T2 #K0 #T0 #H #HT0 #HK0 -elim (lfdeq_fpb_trans … HK12 … H) -K2 #L0 #U0 #H #HUT0 #HLK0 -@(ex2_3_intro … H) -H (**) (* full auto too slow *) -/4 width=3 by ffdeq_intro_dx, lfdeq_trans, tdeq_lfdeq_conf, tdeq_trans/ -qed-. - -(* Inversion lemmas with degree-based equivalence for closures **************) - -(* Basic_2A1: uses: fpb_inv_fleq *) -lemma fpb_inv_ffdeq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → - ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⊥. -#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 -[ #G2 #L2 #T2 #H12 #H elim (ffdeq_inv_gen_sn … H) -H - /3 width=11 by lfdeq_fwd_length, fqu_inv_tdeq/ -| #T2 #_ #HnT #H elim (ffdeq_inv_gen_sn … H) -H /2 width=1 by/ -| #L2 #_ #HnL #H elim (ffdeq_inv_gen_sn … H) -H /2 width=1 by/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpb_lfdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpb_lfdeq.ma deleted file mode 100644 index ee1dea03e..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpb_lfdeq.ma +++ /dev/null @@ -1,50 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfdeq_fqus.ma". -include "basic_2/rt_transition/cpx_lfdeq.ma". -include "basic_2/rt_transition/lpx_lfdeq.ma". -include "basic_2/rt_transition/fpb.ma". - -(* PROPER PARALLEL RST-TRANSITION FOR CLOSURES ******************************) - -(* Properties with degree-based equivalence for local environments **********) - -lemma tdeq_fpb_trans: ∀h,o,U2,U1. U2 ≛[h, o] U1 → - ∀G1,G2,L1,L2,T1. ⦃G1, L1, U1⦄ ≻[h, o] ⦃G2, L2, T1⦄ → - ∃∃L,T2. ⦃G1, L1, U2⦄ ≻[h, o] ⦃G2, L, T2⦄ & T2 ≛[h, o] T1 & L ≛[h, o, T1] L2. -#h #o #U2 #U1 #HU21 #G1 #G2 #L1 #L2 #T1 * -G2 -L2 -T1 -[ #G2 #L2 #T1 #H - elim (tdeq_fqu_trans … H … HU21) -H - /3 width=5 by fpb_fqu, ex3_2_intro/ -| #T1 #HUT1 #HnUT1 - elim (tdeq_cpx_trans … HU21 … HUT1) -HUT1 - /6 width=5 by fpb_cpx, tdeq_canc_sn, tdeq_trans, ex3_2_intro/ -| /6 width=5 by fpb_lpx, lfpx_tdeq_div, tdeq_lfdeq_conf, ex3_2_intro/ -] -qed-. - -(* Basic_2A1: was just: lleq_fpb_trans *) -lemma lfdeq_fpb_trans: ∀h,o,F,K1,K2,T. K1 ≛[h, o, T] K2 → - ∀G,L2,U. ⦃F, K2, T⦄ ≻[h, o] ⦃G, L2, U⦄ → - ∃∃L1,U0. ⦃F, K1, T⦄ ≻[h, o] ⦃G, L1, U0⦄ & U0 ≛[h, o] U & L1 ≛[h, o, U] L2. -#h #o #F #K1 #K2 #T #HT #G #L2 #U * -G -L2 -U -[ #G #L2 #U #H2 elim (lfdeq_fqu_trans … H2 … HT) -K2 - /3 width=5 by fpb_fqu, ex3_2_intro/ -| #U #HTU #HnTU elim (lfdeq_cpx_trans … HT … HTU) -HTU - /5 width=10 by fpb_cpx, cpx_lfdeq_conf_sn, tdeq_trans, tdeq_lfdeq_conf, ex3_2_intro/ -| #L2 #HKL2 #HnKL2 elim (lfdeq_lpx_trans … HKL2 … HT) -HKL2 - /6 width=5 by fpb_lpx, (* 2x *) lfdeq_canc_sn, ex3_2_intro/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpb_rdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpb_rdeq.ma new file mode 100644 index 000000000..6f9a66c72 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpb_rdeq.ma @@ -0,0 +1,50 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rdeq_fqus.ma". +include "basic_2/rt_transition/cpx_rdeq.ma". +include "basic_2/rt_transition/lpx_rdeq.ma". +include "basic_2/rt_transition/fpb.ma". + +(* PROPER PARALLEL RST-TRANSITION FOR CLOSURES ******************************) + +(* Properties with degree-based equivalence for local environments **********) + +lemma tdeq_fpb_trans: ∀h,o,U2,U1. U2 ≛[h, o] U1 → + ∀G1,G2,L1,L2,T1. ⦃G1, L1, U1⦄ ≻[h, o] ⦃G2, L2, T1⦄ → + ∃∃L,T2. ⦃G1, L1, U2⦄ ≻[h, o] ⦃G2, L, T2⦄ & T2 ≛[h, o] T1 & L ≛[h, o, T1] L2. +#h #o #U2 #U1 #HU21 #G1 #G2 #L1 #L2 #T1 * -G2 -L2 -T1 +[ #G2 #L2 #T1 #H + elim (tdeq_fqu_trans … H … HU21) -H + /3 width=5 by fpb_fqu, ex3_2_intro/ +| #T1 #HUT1 #HnUT1 + elim (tdeq_cpx_trans … HU21 … HUT1) -HUT1 + /6 width=5 by fpb_cpx, tdeq_canc_sn, tdeq_trans, ex3_2_intro/ +| /6 width=5 by fpb_lpx, rpx_tdeq_div, tdeq_rdeq_conf, ex3_2_intro/ +] +qed-. + +(* Basic_2A1: was just: lleq_fpb_trans *) +lemma rdeq_fpb_trans: ∀h,o,F,K1,K2,T. K1 ≛[h, o, T] K2 → + ∀G,L2,U. ⦃F, K2, T⦄ ≻[h, o] ⦃G, L2, U⦄ → + ∃∃L1,U0. ⦃F, K1, T⦄ ≻[h, o] ⦃G, L1, U0⦄ & U0 ≛[h, o] U & L1 ≛[h, o, U] L2. +#h #o #F #K1 #K2 #T #HT #G #L2 #U * -G -L2 -U +[ #G #L2 #U #H2 elim (rdeq_fqu_trans … H2 … HT) -K2 + /3 width=5 by fpb_fqu, ex3_2_intro/ +| #U #HTU #HnTU elim (rdeq_cpx_trans … HT … HTU) -HTU + /5 width=10 by fpb_cpx, cpx_rdeq_conf_sn, tdeq_trans, tdeq_rdeq_conf, ex3_2_intro/ +| #L2 #HKL2 #HnKL2 elim (rdeq_lpx_trans … HKL2 … HT) -HKL2 + /6 width=5 by fpb_lpx, (* 2x *) rdeq_canc_sn, ex3_2_intro/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpbq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpbq.ma index 600d9190a..12bd824fa 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpbq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpbq.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "basic_2/notation/relations/predsubty_8.ma". -include "basic_2/static/ffdeq.ma". +include "basic_2/static/fdeq.ma". include "basic_2/s_transition/fquq.ma". include "basic_2/rt_transition/lpr_lpx.ma". @@ -21,10 +21,10 @@ include "basic_2/rt_transition/lpr_lpx.ma". (* Basic_2A1: includes: fleq_fpbq fpbq_lleq *) inductive fpbq (h) (o) (G1) (L1) (T1): relation3 genv lenv term ≝ -| fpbq_fquq : ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → fpbq h o G1 L1 T1 G2 L2 T2 -| fpbq_cpx : ∀T2. ⦃G1, L1⦄ ⊢ T1 ⬈[h] T2 → fpbq h o G1 L1 T1 G1 L1 T2 -| fpbq_lpx : ∀L2. ⦃G1, L1⦄ ⊢ ⬈[h] L2 → fpbq h o G1 L1 T1 G1 L2 T1 -| fpbq_ffdeq: ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → fpbq h o G1 L1 T1 G2 L2 T2 +| fpbq_fquq: ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → fpbq h o G1 L1 T1 G2 L2 T2 +| fpbq_cpx : ∀T2. ⦃G1, L1⦄ ⊢ T1 ⬈[h] T2 → fpbq h o G1 L1 T1 G1 L1 T2 +| fpbq_lpx : ∀L2. ⦃G1, L1⦄ ⊢ ⬈[h] L2 → fpbq h o G1 L1 T1 G1 L2 T1 +| fpbq_fdeq: ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → fpbq h o G1 L1 T1 G2 L2 T2 . interpretation diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpbq_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpbq_aaa.ma index 530e3b8c2..55b946e58 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpbq_aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpbq_aaa.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "basic_2/static/aaa_fqus.ma". -include "basic_2/static/aaa_ffdeq.ma". +include "basic_2/static/aaa_fdeq.ma". include "basic_2/rt_transition/lpx_aaa.ma". include "basic_2/rt_transition/fpbq.ma". @@ -24,5 +24,5 @@ include "basic_2/rt_transition/fpbq.ma". lemma fpbq_aaa_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ → ∀A1. ⦃G1, L1⦄ ⊢ T1 ⁝ A1 → ∃A2. ⦃G2, L2⦄ ⊢ T2 ⁝ A2. #h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 -/3 width=8 by lpx_aaa_conf, cpx_aaa_conf, aaa_ffdeq_conf, aaa_fquq_conf, ex_intro/ +/3 width=8 by lpx_aaa_conf, cpx_aaa_conf, aaa_fdeq_conf, aaa_fquq_conf, ex_intro/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpbq_fpb.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpbq_fpb.ma index 6400d582e..1fa7b7209 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpbq_fpb.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/fpbq_fpb.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/rt_transition/fpb_ffdeq.ma". +include "basic_2/rt_transition/fpb_fdeq.ma". include "basic_2/rt_transition/fpbq.ma". (* PARALLEL RST-TRANSITION FOR CLOSURES *************************************) @@ -28,7 +28,7 @@ qed. (* Basic_2A1: fpb_fpbq_alt *) lemma fpb_fpbq_ffdneq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → ∧∧ ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ & (⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⊥). -/3 width=10 by fpb_fpbq, fpb_inv_ffdeq, conj/ qed-. +/3 width=10 by fpb_fpbq, fpb_inv_fdeq, conj/ qed-. (* Inversrion lemmas with proper parallel rst-transition for closures *******) @@ -38,11 +38,11 @@ lemma fpbq_inv_fpb: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, | ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄. #h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 [ #G2 #L2 #T2 * [2: * #H1 #H2 #H3 destruct ] - /3 width=1 by fpb_fqu, ffdeq_intro_sn, or_intror, or_introl/ + /3 width=1 by fpb_fqu, fdeq_intro_sn, or_intror, or_introl/ | #T2 #H elim (tdeq_dec h o T1 T2) - /4 width=1 by fpb_cpx, ffdeq_intro_sn, or_intror, or_introl/ -| #L2 elim (lfdeq_dec h o L1 L2 T1) - /4 width=1 by fpb_lpx, ffdeq_intro_sn, or_intror, or_introl/ + /4 width=1 by fpb_cpx, fdeq_intro_sn, or_intror, or_introl/ +| #L2 elim (rdeq_dec h o L1 L2 T1) + /4 width=1 by fpb_lpx, fdeq_intro_sn, or_intror, or_introl/ | /2 width=1 by or_introl/ ] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx.ma deleted file mode 100644 index 4067962d8..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx.ma +++ /dev/null @@ -1,137 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/notation/relations/predtysn_5.ma". -include "basic_2/static/lfxs.ma". -include "basic_2/rt_transition/cpx_ext.ma". - -(* UNBOUND PARALLEL RT-TRANSITION FOR REFERRED LOCAL ENVIRONMENTS ***********) - -definition lfpx (h) (G): relation3 term lenv lenv ≝ - lfxs (cpx h G). - -interpretation - "unbound parallel rt-transition on referred entries (local environment)" - 'PRedTySn h T G L1 L2 = (lfpx h G T L1 L2). - -(* Basic properties ***********************************************************) - -lemma lfpx_atom: ∀h,I,G. ⦃G, ⋆⦄ ⊢ ⬈[h, ⓪{I}] ⋆. -/2 width=1 by lfxs_atom/ qed. - -lemma lfpx_sort: ∀h,I1,I2,G,L1,L2,s. - ⦃G, L1⦄ ⊢ ⬈[h, ⋆s] L2 → ⦃G, L1.ⓘ{I1}⦄ ⊢ ⬈[h, ⋆s] L2.ⓘ{I2}. -/2 width=1 by lfxs_sort/ qed. - -lemma lfpx_pair: ∀h,I,G,L1,L2,V1,V2. - ⦃G, L1⦄ ⊢ ⬈[h, V1] L2 → ⦃G, L1⦄ ⊢ V1 ⬈[h] V2 → ⦃G, L1.ⓑ{I}V1⦄ ⊢ ⬈[h, #0] L2.ⓑ{I}V2. -/2 width=1 by lfxs_pair/ qed. - -lemma lfpx_lref: ∀h,I1,I2,G,L1,L2,i. - ⦃G, L1⦄ ⊢ ⬈[h, #i] L2 → ⦃G, L1.ⓘ{I1}⦄ ⊢ ⬈[h, #↑i] L2.ⓘ{I2}. -/2 width=1 by lfxs_lref/ qed. - -lemma lfpx_gref: ∀h,I1,I2,G,L1,L2,l. - ⦃G, L1⦄ ⊢ ⬈[h, §l] L2 → ⦃G, L1.ⓘ{I1}⦄ ⊢ ⬈[h, §l] L2.ⓘ{I2}. -/2 width=1 by lfxs_gref/ qed. - -lemma lfpx_bind_repl_dx: ∀h,I,I1,G,L1,L2,T. - ⦃G, L1.ⓘ{I}⦄ ⊢ ⬈[h, T] L2.ⓘ{I1} → - ∀I2. ⦃G, L1⦄ ⊢ I ⬈[h] I2 → - ⦃G, L1.ⓘ{I}⦄ ⊢ ⬈[h, T] L2.ⓘ{I2}. -/2 width=2 by lfxs_bind_repl_dx/ qed-. - -(* Basic inversion lemmas ***************************************************) - -lemma lfpx_inv_atom_sn: ∀h,G,Y2,T. ⦃G, ⋆⦄ ⊢ ⬈[h, T] Y2 → Y2 = ⋆. -/2 width=3 by lfxs_inv_atom_sn/ qed-. - -lemma lfpx_inv_atom_dx: ∀h,G,Y1,T. ⦃G, Y1⦄ ⊢ ⬈[h, T] ⋆ → Y1 = ⋆. -/2 width=3 by lfxs_inv_atom_dx/ qed-. - -lemma lfpx_inv_sort: ∀h,G,Y1,Y2,s. ⦃G, Y1⦄ ⊢ ⬈[h, ⋆s] Y2 → - ∨∨ Y1 = ⋆ ∧ Y2 = ⋆ - | ∃∃I1,I2,L1,L2. ⦃G, L1⦄ ⊢ ⬈[h, ⋆s] L2 & - Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. -/2 width=1 by lfxs_inv_sort/ qed-. - -lemma lfpx_inv_lref: ∀h,G,Y1,Y2,i. ⦃G, Y1⦄ ⊢ ⬈[h, #↑i] Y2 → - ∨∨ Y1 = ⋆ ∧ Y2 = ⋆ - | ∃∃I1,I2,L1,L2. ⦃G, L1⦄ ⊢ ⬈[h, #i] L2 & - Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. -/2 width=1 by lfxs_inv_lref/ qed-. - -lemma lfpx_inv_gref: ∀h,G,Y1,Y2,l. ⦃G, Y1⦄ ⊢ ⬈[h, §l] Y2 → - ∨∨ Y1 = ⋆ ∧ Y2 = ⋆ - | ∃∃I1,I2,L1,L2. ⦃G, L1⦄ ⊢ ⬈[h, §l] L2 & - Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. -/2 width=1 by lfxs_inv_gref/ qed-. - -lemma lfpx_inv_bind: ∀h,p,I,G,L1,L2,V,T. ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V.T] L2 → - ∧∧ ⦃G, L1⦄ ⊢ ⬈[h, V] L2 & ⦃G, L1.ⓑ{I}V⦄ ⊢ ⬈[h, T] L2.ⓑ{I}V. -/2 width=2 by lfxs_inv_bind/ qed-. - -lemma lfpx_inv_flat: ∀h,I,G,L1,L2,V,T. ⦃G, L1⦄ ⊢ ⬈[h, ⓕ{I}V.T] L2 → - ∧∧ ⦃G, L1⦄ ⊢ ⬈[h, V] L2 & ⦃G, L1⦄ ⊢ ⬈[h, T] L2. -/2 width=2 by lfxs_inv_flat/ qed-. - -(* Advanced inversion lemmas ************************************************) - -lemma lfpx_inv_sort_bind_sn: ∀h,I1,G,Y2,L1,s. ⦃G, L1.ⓘ{I1}⦄ ⊢ ⬈[h, ⋆s] Y2 → - ∃∃I2,L2. ⦃G, L1⦄ ⊢ ⬈[h, ⋆s] L2 & Y2 = L2.ⓘ{I2}. -/2 width=2 by lfxs_inv_sort_bind_sn/ qed-. - -lemma lfpx_inv_sort_bind_dx: ∀h,I2,G,Y1,L2,s. ⦃G, Y1⦄ ⊢ ⬈[h, ⋆s] L2.ⓘ{I2} → - ∃∃I1,L1. ⦃G, L1⦄ ⊢ ⬈[h, ⋆s] L2 & Y1 = L1.ⓘ{I1}. -/2 width=2 by lfxs_inv_sort_bind_dx/ qed-. - -lemma lfpx_inv_zero_pair_sn: ∀h,I,G,Y2,L1,V1. ⦃G, L1.ⓑ{I}V1⦄ ⊢ ⬈[h, #0] Y2 → - ∃∃L2,V2. ⦃G, L1⦄ ⊢ ⬈[h, V1] L2 & ⦃G, L1⦄ ⊢ V1 ⬈[h] V2 & - Y2 = L2.ⓑ{I}V2. -/2 width=1 by lfxs_inv_zero_pair_sn/ qed-. - -lemma lfpx_inv_zero_pair_dx: ∀h,I,G,Y1,L2,V2. ⦃G, Y1⦄ ⊢ ⬈[h, #0] L2.ⓑ{I}V2 → - ∃∃L1,V1. ⦃G, L1⦄ ⊢ ⬈[h, V1] L2 & ⦃G, L1⦄ ⊢ V1 ⬈[h] V2 & - Y1 = L1.ⓑ{I}V1. -/2 width=1 by lfxs_inv_zero_pair_dx/ qed-. - -lemma lfpx_inv_lref_bind_sn: ∀h,I1,G,Y2,L1,i. ⦃G, L1.ⓘ{I1}⦄ ⊢ ⬈[h, #↑i] Y2 → - ∃∃I2,L2. ⦃G, L1⦄ ⊢ ⬈[h, #i] L2 & Y2 = L2.ⓘ{I2}. -/2 width=2 by lfxs_inv_lref_bind_sn/ qed-. - -lemma lfpx_inv_lref_bind_dx: ∀h,I2,G,Y1,L2,i. ⦃G, Y1⦄ ⊢ ⬈[h, #↑i] L2.ⓘ{I2} → - ∃∃I1,L1. ⦃G, L1⦄ ⊢ ⬈[h, #i] L2 & Y1 = L1.ⓘ{I1}. -/2 width=2 by lfxs_inv_lref_bind_dx/ qed-. - -lemma lfpx_inv_gref_bind_sn: ∀h,I1,G,Y2,L1,l. ⦃G, L1.ⓘ{I1}⦄ ⊢ ⬈[h, §l] Y2 → - ∃∃I2,L2. ⦃G, L1⦄ ⊢ ⬈[h, §l] L2 & Y2 = L2.ⓘ{I2}. -/2 width=2 by lfxs_inv_gref_bind_sn/ qed-. - -lemma lfpx_inv_gref_bind_dx: ∀h,I2,G,Y1,L2,l. ⦃G, Y1⦄ ⊢ ⬈[h, §l] L2.ⓘ{I2} → - ∃∃I1,L1. ⦃G, L1⦄ ⊢ ⬈[h, §l] L2 & Y1 = L1.ⓘ{I1}. -/2 width=2 by lfxs_inv_gref_bind_dx/ qed-. - -(* Basic forward lemmas *****************************************************) - -lemma lfpx_fwd_pair_sn: ∀h,I,G,L1,L2,V,T. - ⦃G, L1⦄ ⊢ ⬈[h, ②{I}V.T] L2 → ⦃G, L1⦄ ⊢ ⬈[h, V] L2. -/2 width=3 by lfxs_fwd_pair_sn/ qed-. - -lemma lfpx_fwd_bind_dx: ∀h,p,I,G,L1,L2,V,T. - ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V.T] L2 → ⦃G, L1.ⓑ{I}V⦄ ⊢ ⬈[h, T] L2.ⓑ{I}V. -/2 width=2 by lfxs_fwd_bind_dx/ qed-. - -lemma lfpx_fwd_flat_dx: ∀h,I,G,L1,L2,V,T. - ⦃G, L1⦄ ⊢ ⬈[h, ⓕ{I}V.T] L2 → ⦃G, L1⦄ ⊢ ⬈[h, T] L2. -/2 width=3 by lfxs_fwd_flat_dx/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_fqup.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_fqup.ma deleted file mode 100644 index c3de116d0..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_fqup.ma +++ /dev/null @@ -1,39 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfxs_fqup.ma". -include "basic_2/rt_transition/lfpx.ma". - -(* UNBOUND PARALLEL RT-TRANSITION FOR REFERRED LOCAL ENVIRONMENTS ***********) - -(* Advanced properties ******************************************************) - -lemma lfpx_refl: ∀h,G,T. reflexive … (lfpx h G T). -/2 width=1 by lfxs_refl/ qed. - -lemma lfpx_pair_refl: ∀h,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ⬈[h] V2 → - ∀I,T. ⦃G, L.ⓑ{I}V1⦄ ⊢ ⬈[h, T] L.ⓑ{I}V2. -/2 width=1 by lfxs_pair_refl/ qed. - -(* Advanced inversion lemmas ************************************************) - -lemma lfpx_inv_bind_void: ∀h,p,I,G,L1,L2,V,T. ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V.T] L2 → - ∧∧ ⦃G, L1⦄ ⊢ ⬈[h, V] L2 & ⦃G, L1.ⓧ⦄ ⊢ ⬈[h, T] L2.ⓧ. -/2 width=3 by lfxs_inv_bind_void/ qed-. - -(* Advanced forward lemmas **************************************************) - -lemma lfpx_fwd_bind_dx_void: ∀h,p,I,G,L1,L2,V,T. - ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V.T] L2 → ⦃G, L1.ⓧ⦄ ⊢ ⬈[h, T] L2.ⓧ. -/2 width=4 by lfxs_fwd_bind_dx_void/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_fsle.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_fsle.ma deleted file mode 100644 index 0b7078664..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_fsle.ma +++ /dev/null @@ -1,136 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/fsle_drops.ma". -include "basic_2/static/lfxs_fsle.ma". -include "basic_2/rt_transition/lfpx_length.ma". -include "basic_2/rt_transition/lfpx_fqup.ma". - -(* UNBOUND PARALLEL RT-TRANSITION FOR REFERRED LOCAL ENVIRONMENTS ***********) - -(* Forward lemmas with free variables inclusion for restricted closures *****) - -(* Note: "⦃L2, T1⦄ ⊆ ⦃L2, T0⦄" does not hold *) -(* Note: Take L0 = K0.ⓓ(ⓝW.V), L2 = K0.ⓓW, T0 = #0, T1 = ⬆*[1]V *) -(* Note: This invalidates lfpxs_cpx_conf: "∀h,G. s_r_confluent1 … (cpx h G) (lfpxs h G)" *) -lemma lfpx_cpx_conf_fsge (h) (G): ∀L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ⬈[h] T1 → - ∀L2. ⦃G, L0⦄ ⊢⬈[h, T0] L2 → ⦃L2, T1⦄ ⊆ ⦃L0, T0⦄. -#h #G0 #L0 #T0 @(fqup_wf_ind_eq (Ⓕ) … G0 L0 T0) -G0 -L0 -T0 -#G #L #T #IH #G0 #L0 * * -[ #s #HG #HL #HT #X #HX #Y #HY destruct -IH - elim (cpx_inv_sort1 … HX) -HX #H destruct - lapply (lfpx_fwd_length … HY) -HY #H0 - /2 width=1 by fsle_sort_bi/ -| * [| #i ] #HG #HL #HT #X #HX #Y #HY destruct - [ elim (cpx_inv_zero1 … HX) -HX - [ #H destruct - elim (lfpx_inv_zero_length … HY) -HY * - [ #H1 #H2 destruct -IH // - | #I #K0 #K2 #V0 #V2 #HK02 #HV02 #H1 #H2 destruct - lapply (lfpx_fwd_length … HK02) #H0 - /4 width=4 by fsle_pair_bi, fqu_fqup, fqu_lref_O/ - | #I #K0 #K2 #HK02 #H1 #H2 destruct -IH - /2 width=1 by fsle_unit_bi/ - ] - | * #I0 #K0 #V0 #V1 #HV01 #HV1X #H destruct - elim (lfpx_inv_zero_pair_sn … HY) -HY #K2 #V2 #HK02 #HV02 #H destruct - lapply (lfpx_fwd_length … HK02) #H0 - /4 width=4 by fsle_lifts_SO_sn, fqu_fqup, fqu_lref_O/ - ] - | elim (cpx_inv_lref1 … HX) -HX - [ #H destruct - elim (lfpx_inv_lref … HY) -HY * - [ #H0 #H1 destruct // - | #I0 #I2 #K0 #K2 #HK02 #H1 #H2 destruct - lapply (lfpx_fwd_length … HK02) #H0 - /4 width=5 by fsle_lifts_SO, fqu_fqup/ - ] - | * #I0 #K0 #V1 #HV1 #HV1X #H0 destruct - elim (lfpx_inv_lref_bind_sn … HY) -HY #I2 #K2 #HK02 #H destruct - lapply (lfpx_fwd_length … HK02) #H0 - /4 width=5 by fsle_lifts_SO, fqu_fqup/ - ] - ] -| #l #HG #HL #HT #X #HX #Y #HY destruct -IH - >(cpx_inv_gref1 … HX) -X - lapply (lfpx_fwd_length … HY) -HY #H0 - /2 width=1 by fsle_gref_bi/ -| #p #I #V0 #T0 #HG #HL #HT #X #HX #Y #HY destruct - elim (lfpx_inv_bind … V0 ? HY) -HY #HV0 #HT0 - elim (cpx_inv_bind1 … HX) -HX * - [ #V1 #T1 #HV01 #HT01 #H destruct - lapply (lfpx_fwd_length … HV0) #H0 - /4 width=6 by fsle_bind_eq, fsle_fwd_pair_sn/ - | #T #HT #HXT #H1 #H2 destruct - lapply (lfpx_fwd_length … HV0) #H0 - /3 width=8 by fsle_inv_lifts_sn/ - ] -| #I #V0 #X0 #HG #HL #HT #X #HX #Y #HY destruct - elim (lfxs_inv_flat … HY) -HY #HV0 #HX0 - elim (cpx_inv_flat1 … HX) -HX * - [ #V1 #T1 #HV01 #HT01 #H destruct - /3 width=4 by fsle_flat/ - | #HX #H destruct - /4 width=4 by fsle_flat_dx_dx/ - | #HX #H destruct - /4 width=4 by fsle_flat_dx_sn/ - | #p #V1 #W0 #W1 #T0 #T1 #HV01 #HW01 #HT01 #H1 #H2 #H3 destruct - elim (lfpx_inv_bind … W0 ? HX0) -HX0 #HW0 #HT0 - lapply (lfpx_fwd_length … HV0) #H0 - lapply (IH … HV01 … HV0) -HV01 -HV0 // #HV - lapply (IH … HW01 … HW0) -HW01 -HW0 // #HW - lapply (IH … HT01 … HT0) -HT01 -HT0 -IH // #HT - lapply (fsle_fwd_pair_sn … HT) -HT #HT - @fsle_bind_sn_ge // - [ /4 width=1 by fsle_flat_sn, fsle_flat_dx_dx, fsle_flat_dx_sn, fsle_bind_dx_sn/ - | /3 width=1 by fsle_flat_dx_dx, fsle_shift/ - ] - | #p #V1 #X1 #W0 #W1 #T0 #T1 #HV01 #HVX1 #HW01 #HT01 #H1 #H2 #H3 destruct - elim (lfpx_inv_bind … W0 ? HX0) -HX0 #HW0 #HT0 - lapply (lfpx_fwd_length … HV0) #H0 - lapply (IH … HV01 … HV0) -HV01 -HV0 // #HV - lapply (IH … HW01 … HW0) -HW01 -HW0 // #HW - lapply (IH … HT01 … HT0) -HT01 -HT0 -IH // #HT - lapply (fsle_fwd_pair_sn … HT) -HT #HT - @fsle_bind_sn_ge // - [ /3 width=1 by fsle_flat_dx_dx, fsle_bind_dx_sn/ - | /4 width=3 by fsle_flat_sn, fsle_flat_dx_sn, fsle_flat_dx_dx, fsle_shift, fsle_lifts_sn/ - ] - ] -] -qed-. - -lemma lfpx_fsge_comp (h) (G): lfxs_fsge_compatible (cpx h G). -/2 width=4 by lfpx_cpx_conf_fsge/ qed-. - -(**) (* this section concerns cpx *) -(* Properties with generic extension on referred entries ********************) - -(* Basic_2A1: uses: cpx_frees_trans *) -lemma cpx_fsge_comp (h) (G): R_fsge_compatible (cpx h G). -/2 width=4 by lfpx_cpx_conf_fsge/ qed-. - -(* Note: lemma 1000 *) -(* Basic_2A1: uses: cpx_llpx_sn_conf *) -lemma cpx_lfxs_conf (R) (h) (G): s_r_confluent1 … (cpx h G) (lfxs R). -/3 width=3 by fsge_lfxs_trans, cpx_fsge_comp/ qed-. - -(* Advanced properties ******************************************************) - -lemma lfpx_cpx_conf (h) (G): s_r_confluent1 … (cpx h G) (lfpx h G). -/2 width=5 by cpx_lfxs_conf/ qed-. - -lemma lfpx_cpx_conf_fsge_dx (h) (G): ∀L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ⬈[h] T1 → - ∀L2. ⦃G, L0⦄ ⊢⬈[h, T0] L2 → ⦃L2, T1⦄ ⊆ ⦃L0, T1⦄. -/3 width=5 by lfpx_cpx_conf, lfpx_fsge_comp/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_length.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_length.ma deleted file mode 100644 index e88d45eef..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_length.ma +++ /dev/null @@ -1,33 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfxs_length.ma". -include "basic_2/rt_transition/lfpx.ma". - -(* UNBOUND PARALLEL RT-TRANSITION FOR REFERRED LOCAL ENVIRONMENTS ***********) - -(* Forward lemmas with length for local environments ************************) - -lemma lfpx_fwd_length: ∀h,G,L1,L2,T. ⦃G, L1⦄ ⊢ ⬈[h, T] L2 → |L1| = |L2|. -/2 width=3 by lfxs_fwd_length/ qed-. - -(* Inversion lemmas with length for local environments **********************) - -lemma lfpx_inv_zero_length: ∀h,G,Y1,Y2. ⦃G, Y1⦄ ⊢ ⬈[h, #0] Y2 → - ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆ - | ∃∃I,L1,L2,V1,V2. ⦃G, L1⦄ ⊢ ⬈[h, V1] L2 & - ⦃G, L1⦄ ⊢ V1 ⬈[h] V2 & - Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2 - |∃∃I,L1,L2. |L1| = |L2| & Y1 = L1.ⓤ{I} & Y2 = L2.ⓤ{I}. -/2 width=1 by lfxs_inv_zero_length/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma deleted file mode 100644 index 668246dd1..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma +++ /dev/null @@ -1,168 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfdeq_fqup.ma". -include "basic_2/static/lfdeq_lfdeq.ma". -include "basic_2/rt_transition/lfpx_fsle.ma". - -(* UNBOUND PARALLEL RT-TRANSITION FOR REFERRED LOCAL ENVIRONMENTS ***********) - -(* Properties with degree-based equivalence for local environments **********) - -lemma lfpx_pair_sn_split: ∀h,G,L1,L2,V. ⦃G, L1⦄ ⊢ ⬈[h, V] L2 → ∀o,I,T. - ∃∃L. ⦃G, L1⦄ ⊢ ⬈[h, ②{I}V.T] L & L ≛[h, o, V] L2. -/3 width=5 by lfpx_fsge_comp, lfxs_pair_sn_split/ qed-. - -lemma lfpx_flat_dx_split: ∀h,G,L1,L2,T. ⦃G, L1⦄ ⊢ ⬈[h, T] L2 → ∀o,I,V. - ∃∃L. ⦃G, L1⦄ ⊢ ⬈[h, ⓕ{I}V.T] L & L ≛[h, o, T] L2. -/3 width=5 by lfpx_fsge_comp, lfxs_flat_dx_split/ qed-. - -lemma lfpx_bind_dx_split: ∀h,I,G,L1,L2,V1,T. ⦃G, L1.ⓑ{I}V1⦄ ⊢ ⬈[h, T] L2 → ∀o,p. - ∃∃L,V. ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V1.T] L & L.ⓑ{I}V ≛[h, o, T] L2 & ⦃G, L1⦄ ⊢ V1 ⬈[h] V. -/3 width=5 by lfpx_fsge_comp, lfxs_bind_dx_split/ qed-. - -lemma lfpx_bind_dx_split_void: ∀h,G,K1,L2,T. ⦃G, K1.ⓧ⦄ ⊢ ⬈[h, T] L2 → ∀o,p,I,V. - ∃∃K2. ⦃G, K1⦄ ⊢ ⬈[h, ⓑ{p,I}V.T] K2 & K2.ⓧ ≛[h, o, T] L2. -/3 width=5 by lfpx_fsge_comp, lfxs_bind_dx_split_void/ qed-. - -lemma lfpx_tdeq_conf: ∀h,o,G. s_r_confluent1 … (cdeq h o) (lfpx h G). -/2 width=5 by tdeq_lfxs_conf/ qed-. - -lemma lfpx_tdeq_div: ∀h,o,T1,T2. T1 ≛[h, o] T2 → - ∀G,L1,L2. ⦃G, L1⦄ ⊢ ⬈[h, T2] L2 → ⦃G, L1⦄ ⊢ ⬈[h, T1] L2. -/2 width=5 by tdeq_lfxs_div/ qed-. - -lemma cpx_tdeq_conf_lexs: ∀h,o,G. R_confluent2_lfxs … (cpx h G) (cdeq h o) (cpx h G) (cdeq h o). -#h #o #G #L0 #T0 #T1 #H @(cpx_ind … H) -G -L0 -T0 -T1 /2 width=3 by ex2_intro/ -[ #G #L0 #s0 #X0 #H0 #L1 #HL01 #L2 #HL02 - elim (tdeq_inv_sort1 … H0) -H0 #s1 #d1 #Hs0 #Hs1 #H destruct - /4 width=3 by tdeq_sort, deg_next, ex2_intro/ -| #I #G #K0 #V0 #V1 #W1 #_ #IH #HVW1 #T2 #H0 #L1 #H1 #L2 #H2 - >(tdeq_inv_lref1 … H0) -H0 - elim (lfpx_inv_zero_pair_sn … H1) -H1 #K1 #X1 #HK01 #HX1 #H destruct - elim (lfdeq_inv_zero_pair_sn … H2) -H2 #K2 #X2 #HK02 #HX2 #H destruct - elim (IH X2 … HK01 … HK02) // -K0 -V0 #V #HV1 #HV2 - elim (tdeq_lifts_sn … HV1 … HVW1) -V1 /3 width=5 by cpx_delta, ex2_intro/ -| #I0 #G #K0 #V1 #W1 #i #_ #IH #HVW1 #T2 #H0 #L1 #H1 #L2 #H2 - >(tdeq_inv_lref1 … H0) -H0 - elim (lfpx_inv_lref_bind_sn … H1) -H1 #I1 #K1 #HK01 #H destruct - elim (lfdeq_inv_lref_bind_sn … H2) -H2 #I2 #K2 #HK02 #H destruct - elim (IH … HK01 … HK02) [|*: //] -K0 #V #HV1 #HV2 - elim (tdeq_lifts_sn … HV1 … HVW1) -V1 /3 width=5 by cpx_lref, ex2_intro/ -| #p #I #G #L0 #V0 #V1 #T0 #T1 #_ #_ #IHV #IHT #X0 #H0 #L1 #H1 #L2 #H2 - elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #HV02 #HT02 #H destruct - elim (lfpx_inv_bind … H1) -H1 #HL01 #H1 - elim (lfdeq_inv_bind … H2) -H2 #HL02 #H2 - lapply (lfdeq_bind_repl_dx … H2 (BPair I V2) ?) -H2 /2 width=1 by ext2_pair/ #H2 - elim (IHV … HV02 … HL01 … HL02) -IHV -HV02 -HL01 -HL02 - elim (IHT … HT02 … H1 … H2) -L0 -T0 - /3 width=5 by cpx_bind, tdeq_pair, ex2_intro/ -| #I #G #L0 #V0 #V1 #T0 #T1 #_ #_ #IHV #IHT #X0 #H0 #L1 #H1 #L2 #H2 - elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #HV02 #HT02 #H destruct - elim (lfpx_inv_flat … H1) -H1 #HL01 #H1 - elim (lfdeq_inv_flat … H2) -H2 #HL02 #H2 - elim (IHV … HV02 … HL01 … HL02) -IHV -HV02 -HL01 -HL02 - elim (IHT … HT02 … H1 … H2) -L0 -V0 -T0 - /3 width=5 by cpx_flat, tdeq_pair, ex2_intro/ -| #G #L0 #V0 #T0 #T1 #U1 #_ #IH #HUT1 #X0 #H0 #L1 #H1 #L2 #H2 - elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #HV02 #HT02 #H destruct - elim (lfpx_inv_bind … H1) -H1 #HL01 #H1 - elim (lfdeq_inv_bind … H2) -H2 #HL02 #H2 - lapply (lfdeq_bind_repl_dx … H2 (BPair Abbr V2) ?) -H2 /2 width=1 by ext2_pair/ -HV02 #H2 - elim (IH … HT02 … H1 … H2) -L0 -T0 #T #HT1 - elim (tdeq_inv_lifts_sn … HT1 … HUT1) -T1 - /3 width=5 by cpx_zeta, ex2_intro/ -| #G #L0 #V0 #T0 #T1 #_ #IH #X0 #H0 #L1 #H1 #L2 #H2 - elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #_ #HT02 #H destruct - elim (lfpx_inv_flat … H1) -H1 #HL01 #H1 - elim (lfdeq_inv_flat … H2) -H2 #HL02 #H2 - elim (IH … HT02 … H1 … H2) -L0 -V0 -T0 - /3 width=3 by cpx_eps, ex2_intro/ -| #G #L0 #V0 #T0 #T1 #_ #IH #X0 #H0 #L1 #H1 #L2 #H2 - elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #HV02 #_ #H destruct - elim (lfpx_inv_flat … H1) -H1 #HL01 #H1 - elim (lfdeq_inv_flat … H2) -H2 #HL02 #H2 - elim (IH … HV02 … HL01 … HL02) -L0 -V0 -T1 - /3 width=3 by cpx_ee, ex2_intro/ -| #p #G #L0 #V0 #V1 #W0 #W1 #T0 #T1 #_ #_ #_ #IHV #IHW #IHT #X0 #H0 #L1 #H1 #L2 #H2 - elim (tdeq_inv_pair1 … H0) -H0 #V2 #X #HV02 #H0 #H destruct - elim (tdeq_inv_pair1 … H0) -H0 #W2 #T2 #HW02 #HT02 #H destruct - elim (lfpx_inv_flat … H1) -H1 #H1LV0 #H1 - elim (lfpx_inv_bind … H1) -H1 #H1LW0 #H1LT0 - elim (lfdeq_inv_flat … H2) -H2 #H2LV0 #H2 - elim (lfdeq_inv_bind … H2) -H2 #H2LW0 #H2LT0 - lapply (lfdeq_bind_repl_dx … H2LT0 (BPair Abst W2) ?) -H2LT0 /2 width=1 by ext2_pair/ #H2LT0 - elim (IHV … HV02 … H1LV0 … H2LV0) -IHV -HV02 -H1LV0 -H2LV0 - elim (IHW … HW02 … H1LW0 … H2LW0) -IHW -HW02 -H1LW0 -H2LW0 - elim (IHT … HT02 … H1LT0 … H2LT0) -L0 -V0 -T0 - /4 width=7 by cpx_beta, tdeq_pair, ex2_intro/ (* note: 2 tdeq_pair *) -| #p #G #L0 #V0 #V1 #U1 #W0 #W1 #T0 #T1 #_ #_ #_ #IHV #IHW #IHT #HVU1 #X0 #H0 #L1 #H1 #L2 #H2 - elim (tdeq_inv_pair1 … H0) -H0 #V2 #X #HV02 #H0 #H destruct - elim (tdeq_inv_pair1 … H0) -H0 #W2 #T2 #HW02 #HT02 #H destruct - elim (lfpx_inv_flat … H1) -H1 #H1LV0 #H1 - elim (lfpx_inv_bind … H1) -H1 #H1LW0 #H1LT0 - elim (lfdeq_inv_flat … H2) -H2 #H2LV0 #H2 - elim (lfdeq_inv_bind … H2) -H2 #H2LW0 #H2LT0 - lapply (lfdeq_bind_repl_dx … H2LT0 (BPair Abbr W2) ?) -H2LT0 /2 width=1 by ext2_pair/ #H2LT0 - elim (IHV … HV02 … H1LV0 … H2LV0) -IHV -HV02 -H1LV0 -H2LV0 #V #HV1 - elim (IHW … HW02 … H1LW0 … H2LW0) -IHW -HW02 -H1LW0 -H2LW0 - elim (IHT … HT02 … H1LT0 … H2LT0) -L0 -V0 -T0 - elim (tdeq_lifts_sn … HV1 … HVU1) -V1 - /4 width=9 by cpx_theta, tdeq_pair, ex2_intro/ (* note: 2 tdeq_pair *) -] -qed-. - -lemma cpx_tdeq_conf: ∀h,o,G,L. ∀T0:term. ∀T1. ⦃G, L⦄ ⊢ T0 ⬈[h] T1 → - ∀T2. T0 ≛[h, o] T2 → - ∃∃T. T1 ≛[h, o] T & ⦃G, L⦄ ⊢ T2 ⬈[h] T. -#h #o #G #L #T0 #T1 #HT01 #T2 #HT02 -elim (cpx_tdeq_conf_lexs … HT01 … HT02 L … L) -HT01 -HT02 -/2 width=3 by lfxs_refl, ex2_intro/ -qed-. - -lemma tdeq_cpx_trans: ∀h,o,G,L,T2. ∀T0:term. T2 ≛[h, o] T0 → - ∀T1. ⦃G, L⦄ ⊢ T0 ⬈[h] T1 → - ∃∃T. ⦃G, L⦄ ⊢ T2 ⬈[h] T & T ≛[h, o] T1. -#h #o #G #L #T2 #T0 #HT20 #T1 #HT01 -elim (cpx_tdeq_conf … HT01 T2) -HT01 /3 width=3 by tdeq_sym, ex2_intro/ -qed-. - -(* Basic_2A1: uses: cpx_lleq_conf *) -lemma cpx_lfdeq_conf: ∀h,o,G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ⬈[h] T1 → - ∀L2. L0 ≛[h, o, T0] L2 → - ∃∃T. ⦃G, L2⦄ ⊢ T0 ⬈[h] T & T1 ≛[h, o] T. -#h #o #G #L0 #T0 #T1 #HT01 #L2 #HL02 -elim (cpx_tdeq_conf_lexs … HT01 T0 … L0 … HL02) -HT01 -HL02 -/2 width=3 by lfxs_refl, ex2_intro/ -qed-. - -(* Basic_2A1: uses: lleq_cpx_trans *) -lemma lfdeq_cpx_trans: ∀h,o,G,L2,L0,T0. L2 ≛[h, o, T0] L0 → - ∀T1. ⦃G, L0⦄ ⊢ T0 ⬈[h] T1 → - ∃∃T. ⦃G, L2⦄ ⊢ T0 ⬈[h] T & T ≛[h, o] T1. -#h #o #G #L2 #L0 #T0 #HL20 #T1 #HT01 -elim (cpx_lfdeq_conf … o … HT01 L2) -HT01 -/3 width=3 by lfdeq_sym, tdeq_sym, ex2_intro/ -qed-. - -lemma lfpx_lfdeq_conf: ∀h,o,G,T. confluent2 … (lfpx h G T) (lfdeq h o T). -/3 width=6 by lfpx_fsge_comp, lfdeq_fsge_comp, cpx_tdeq_conf_lexs, lfxs_conf/ qed-. - -lemma lfdeq_lfpx_trans: ∀h,o,G,T,L2,K2. ⦃G, L2⦄ ⊢ ⬈[h, T] K2 → - ∀L1. L1 ≛[h, o, T] L2 → - ∃∃K1. ⦃G, L1⦄ ⊢ ⬈[h, T] K1 & K1 ≛[h, o, T] K2. -#h #o #G #T #L2 #K2 #HLK2 #L1 #HL12 -elim (lfpx_lfdeq_conf … o … HLK2 L1) -/3 width=3 by lfdeq_sym, ex2_intro/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfpx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfpx.ma deleted file mode 100644 index a96910a0c..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfpx.ma +++ /dev/null @@ -1,34 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfxs_lfxs.ma". -include "basic_2/rt_transition/lfpx.ma". - -(* UNBOUND PARALLEL RT-TRANSITION FOR REFERRED LOCAL ENVIRONMENTS ***********) - -(* Main properties **********************************************************) - -theorem lfpx_bind: ∀h,G,L1,L2,V1. ⦃G, L1⦄ ⊢ ⬈[h, V1] L2 → - ∀I,V2,T. ⦃G, L1.ⓑ{I}V1⦄ ⊢ ⬈[h, T] L2.ⓑ{I}V2 → - ∀p. ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V1.T] L2. -/2 width=2 by lfxs_bind/ qed. - -theorem lfpx_flat: ∀h,G,L1,L2,V. ⦃G, L1⦄ ⊢ ⬈[h, V] L2 → - ∀I,T. ⦃G, L1⦄ ⊢ ⬈[h, T] L2 → ⦃G, L1⦄ ⊢ ⬈[h, ⓕ{I}V.T] L2. -/2 width=1 by lfxs_flat/ qed. - -theorem lfpx_bind_void: ∀h,G,L1,L2,V. ⦃G, L1⦄ ⊢ ⬈[h, V] L2 → - ∀T. ⦃G, L1.ⓧ⦄ ⊢ ⬈[h, T] L2.ⓧ → - ∀p,I. ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V.T] L2. -/2 width=1 by lfxs_bind_void/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lpx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lpx.ma deleted file mode 100644 index 7ca9a2f28..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lpx.ma +++ /dev/null @@ -1,35 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfxs_lex.ma". -include "basic_2/rt_transition/lfpx_fsle.ma". -include "basic_2/rt_transition/lpx.ma". - -(* UNBOUND PARALLEL RT-TRANSITION FOR REFERRED LOCAL ENVIRONMENTS ***********) - -(* Properties with syntactic equivalence for referred local environments ****) - -lemma fleq_lfpx (h) (G): ∀L1,L2,T. L1 ≡[T] L2 → ⦃G, L1⦄ ⊢ ⬈[h, T] L2. -/2 width=1 by lfeq_fwd_lfxs/ qed. - -(* Properties with unbound parallel rt-transition for full local envs *******) - -lemma lpx_lfpx: ∀h,G,L1,L2,T. ⦃G, L1⦄ ⊢ ⬈[h] L2 → ⦃G, L1⦄ ⊢ ⬈[h, T] L2. -/2 width=1 by lfxs_lex/ qed. - -(* Inversion lemmas with unbound parallel rt-transition for full local envs *) - -lemma lfpx_inv_lpx_lfeq: ∀h,G,L1,L2,T. ⦃G, L1⦄ ⊢ ⬈[h, T] L2 → - ∃∃L. ⦃G, L1⦄ ⊢ ⬈[h] L & L ≡[T] L2. -/3 width=3 by lfpx_fsge_comp, lfxs_inv_lex_lfeq/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lpx_fsle.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lpx_fsle.ma index 7773a081e..f40b6b627 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lpx_fsle.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lpx_fsle.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/rt_transition/lfpx_lpx.ma". +include "basic_2/rt_transition/rpx_lpx.ma". (* UNBOUND PARALLEL RT-TRANSITION FOR FULL LOCAL ENVIRONMENTS ***************) @@ -21,7 +21,7 @@ include "basic_2/rt_transition/lfpx_lpx.ma". (* Basic_2A1: uses: lpx_cpx_frees_trans *) lemma lpx_cpx_conf_fsge (h) (G): ∀L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ⬈[h] T1 → ∀L2. ⦃G, L0⦄ ⊢ ⬈[h] L2 → ⦃L2, T1⦄ ⊆ ⦃L0, T0⦄. -/3 width=4 by lfpx_cpx_conf_fsge, lpx_lfpx/ qed-. +/3 width=4 by rpx_cpx_conf_fsge, lpx_rpx/ qed-. (* Basic_2A1: uses: lpx_frees_trans *) lemma lpx_fsge_comp (h) (G): ∀L0,L2,T0. ⦃G, L0⦄ ⊢ ⬈[h] L2 → ⦃L2, T0⦄ ⊆ ⦃L0, T0⦄. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lpx_lfdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lpx_lfdeq.ma deleted file mode 100644 index 0008d2f79..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lpx_lfdeq.ma +++ /dev/null @@ -1,32 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfdeq_lfeq.ma". -include "basic_2/rt_transition/lfpx_lfdeq.ma". -include "basic_2/rt_transition/lfpx_lpx.ma". - -(* UNBOUND PARALLEL RT-TRANSITION FOR FULL LOCAL ENVIRONMENTS ***************) - -(* Properties with degree-based equivalence for local environments **********) - -(* Basic_2A1: uses: lleq_lpx_trans *) -lemma lfdeq_lpx_trans (h) (o) (G): ∀L2,K2. ⦃G, L2⦄ ⊢ ⬈[h] K2 → - ∀L1. ∀T:term. L1 ≛[h, o, T] L2 → - ∃∃K1. ⦃G, L1⦄ ⊢ ⬈[h] K1 & K1 ≛[h, o, T] K2. -#h #o #G #L2 #K2 #HLK2 #L1 #T #HL12 -lapply (lpx_lfpx … T HLK2) -HLK2 #HLK2 -elim (lfdeq_lfpx_trans … HLK2 … HL12) -L2 #K #H #HK2 -elim (lfpx_inv_lpx_lfeq … H) -H #K1 #HLK1 #HK1 -/3 width=5 by lfeq_lfdeq_trans, ex2_intro/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lpx_rdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lpx_rdeq.ma new file mode 100644 index 000000000..f63af0dcf --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lpx_rdeq.ma @@ -0,0 +1,32 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rdeq_req.ma". +include "basic_2/rt_transition/rpx_rdeq.ma". +include "basic_2/rt_transition/rpx_lpx.ma". + +(* UNBOUND PARALLEL RT-TRANSITION FOR FULL LOCAL ENVIRONMENTS ***************) + +(* Properties with degree-based equivalence for local environments **********) + +(* Basic_2A1: uses: lleq_lpx_trans *) +lemma rdeq_lpx_trans (h) (o) (G): ∀L2,K2. ⦃G, L2⦄ ⊢ ⬈[h] K2 → + ∀L1. ∀T:term. L1 ≛[h, o, T] L2 → + ∃∃K1. ⦃G, L1⦄ ⊢ ⬈[h] K1 & K1 ≛[h, o, T] K2. +#h #o #G #L2 #K2 #HLK2 #L1 #T #HL12 +lapply (lpx_rpx … T HLK2) -HLK2 #HLK2 +elim (rdeq_rpx_trans … HLK2 … HL12) -L2 #K #H #HK2 +elim (rpx_inv_lpx_req … H) -H #K1 #HLK1 #HK1 +/3 width=5 by req_rdeq_trans, ex2_intro/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx.ma new file mode 100644 index 000000000..76364b2f1 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx.ma @@ -0,0 +1,137 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/notation/relations/predtysn_5.ma". +include "basic_2/static/rex.ma". +include "basic_2/rt_transition/cpx_ext.ma". + +(* UNBOUND PARALLEL RT-TRANSITION FOR REFERRED LOCAL ENVIRONMENTS ***********) + +definition rpx (h) (G): relation3 term lenv lenv ≝ + rex (cpx h G). + +interpretation + "unbound parallel rt-transition on referred entries (local environment)" + 'PRedTySn h T G L1 L2 = (rpx h G T L1 L2). + +(* Basic properties ***********************************************************) + +lemma rpx_atom: ∀h,I,G. ⦃G, ⋆⦄ ⊢ ⬈[h, ⓪{I}] ⋆. +/2 width=1 by rex_atom/ qed. + +lemma rpx_sort: ∀h,I1,I2,G,L1,L2,s. + ⦃G, L1⦄ ⊢ ⬈[h, ⋆s] L2 → ⦃G, L1.ⓘ{I1}⦄ ⊢ ⬈[h, ⋆s] L2.ⓘ{I2}. +/2 width=1 by rex_sort/ qed. + +lemma rpx_pair: ∀h,I,G,L1,L2,V1,V2. + ⦃G, L1⦄ ⊢ ⬈[h, V1] L2 → ⦃G, L1⦄ ⊢ V1 ⬈[h] V2 → ⦃G, L1.ⓑ{I}V1⦄ ⊢ ⬈[h, #0] L2.ⓑ{I}V2. +/2 width=1 by rex_pair/ qed. + +lemma rpx_lref: ∀h,I1,I2,G,L1,L2,i. + ⦃G, L1⦄ ⊢ ⬈[h, #i] L2 → ⦃G, L1.ⓘ{I1}⦄ ⊢ ⬈[h, #↑i] L2.ⓘ{I2}. +/2 width=1 by rex_lref/ qed. + +lemma rpx_gref: ∀h,I1,I2,G,L1,L2,l. + ⦃G, L1⦄ ⊢ ⬈[h, §l] L2 → ⦃G, L1.ⓘ{I1}⦄ ⊢ ⬈[h, §l] L2.ⓘ{I2}. +/2 width=1 by rex_gref/ qed. + +lemma rpx_bind_repl_dx: ∀h,I,I1,G,L1,L2,T. + ⦃G, L1.ⓘ{I}⦄ ⊢ ⬈[h, T] L2.ⓘ{I1} → + ∀I2. ⦃G, L1⦄ ⊢ I ⬈[h] I2 → + ⦃G, L1.ⓘ{I}⦄ ⊢ ⬈[h, T] L2.ⓘ{I2}. +/2 width=2 by rex_bind_repl_dx/ qed-. + +(* Basic inversion lemmas ***************************************************) + +lemma rpx_inv_atom_sn: ∀h,G,Y2,T. ⦃G, ⋆⦄ ⊢ ⬈[h, T] Y2 → Y2 = ⋆. +/2 width=3 by rex_inv_atom_sn/ qed-. + +lemma rpx_inv_atom_dx: ∀h,G,Y1,T. ⦃G, Y1⦄ ⊢ ⬈[h, T] ⋆ → Y1 = ⋆. +/2 width=3 by rex_inv_atom_dx/ qed-. + +lemma rpx_inv_sort: ∀h,G,Y1,Y2,s. ⦃G, Y1⦄ ⊢ ⬈[h, ⋆s] Y2 → + ∨∨ Y1 = ⋆ ∧ Y2 = ⋆ + | ∃∃I1,I2,L1,L2. ⦃G, L1⦄ ⊢ ⬈[h, ⋆s] L2 & + Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. +/2 width=1 by rex_inv_sort/ qed-. + +lemma rpx_inv_lref: ∀h,G,Y1,Y2,i. ⦃G, Y1⦄ ⊢ ⬈[h, #↑i] Y2 → + ∨∨ Y1 = ⋆ ∧ Y2 = ⋆ + | ∃∃I1,I2,L1,L2. ⦃G, L1⦄ ⊢ ⬈[h, #i] L2 & + Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. +/2 width=1 by rex_inv_lref/ qed-. + +lemma rpx_inv_gref: ∀h,G,Y1,Y2,l. ⦃G, Y1⦄ ⊢ ⬈[h, §l] Y2 → + ∨∨ Y1 = ⋆ ∧ Y2 = ⋆ + | ∃∃I1,I2,L1,L2. ⦃G, L1⦄ ⊢ ⬈[h, §l] L2 & + Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. +/2 width=1 by rex_inv_gref/ qed-. + +lemma rpx_inv_bind: ∀h,p,I,G,L1,L2,V,T. ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V.T] L2 → + ∧∧ ⦃G, L1⦄ ⊢ ⬈[h, V] L2 & ⦃G, L1.ⓑ{I}V⦄ ⊢ ⬈[h, T] L2.ⓑ{I}V. +/2 width=2 by rex_inv_bind/ qed-. + +lemma rpx_inv_flat: ∀h,I,G,L1,L2,V,T. ⦃G, L1⦄ ⊢ ⬈[h, ⓕ{I}V.T] L2 → + ∧∧ ⦃G, L1⦄ ⊢ ⬈[h, V] L2 & ⦃G, L1⦄ ⊢ ⬈[h, T] L2. +/2 width=2 by rex_inv_flat/ qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma rpx_inv_sort_bind_sn: ∀h,I1,G,Y2,L1,s. ⦃G, L1.ⓘ{I1}⦄ ⊢ ⬈[h, ⋆s] Y2 → + ∃∃I2,L2. ⦃G, L1⦄ ⊢ ⬈[h, ⋆s] L2 & Y2 = L2.ⓘ{I2}. +/2 width=2 by rex_inv_sort_bind_sn/ qed-. + +lemma rpx_inv_sort_bind_dx: ∀h,I2,G,Y1,L2,s. ⦃G, Y1⦄ ⊢ ⬈[h, ⋆s] L2.ⓘ{I2} → + ∃∃I1,L1. ⦃G, L1⦄ ⊢ ⬈[h, ⋆s] L2 & Y1 = L1.ⓘ{I1}. +/2 width=2 by rex_inv_sort_bind_dx/ qed-. + +lemma rpx_inv_zero_pair_sn: ∀h,I,G,Y2,L1,V1. ⦃G, L1.ⓑ{I}V1⦄ ⊢ ⬈[h, #0] Y2 → + ∃∃L2,V2. ⦃G, L1⦄ ⊢ ⬈[h, V1] L2 & ⦃G, L1⦄ ⊢ V1 ⬈[h] V2 & + Y2 = L2.ⓑ{I}V2. +/2 width=1 by rex_inv_zero_pair_sn/ qed-. + +lemma rpx_inv_zero_pair_dx: ∀h,I,G,Y1,L2,V2. ⦃G, Y1⦄ ⊢ ⬈[h, #0] L2.ⓑ{I}V2 → + ∃∃L1,V1. ⦃G, L1⦄ ⊢ ⬈[h, V1] L2 & ⦃G, L1⦄ ⊢ V1 ⬈[h] V2 & + Y1 = L1.ⓑ{I}V1. +/2 width=1 by rex_inv_zero_pair_dx/ qed-. + +lemma rpx_inv_lref_bind_sn: ∀h,I1,G,Y2,L1,i. ⦃G, L1.ⓘ{I1}⦄ ⊢ ⬈[h, #↑i] Y2 → + ∃∃I2,L2. ⦃G, L1⦄ ⊢ ⬈[h, #i] L2 & Y2 = L2.ⓘ{I2}. +/2 width=2 by rex_inv_lref_bind_sn/ qed-. + +lemma rpx_inv_lref_bind_dx: ∀h,I2,G,Y1,L2,i. ⦃G, Y1⦄ ⊢ ⬈[h, #↑i] L2.ⓘ{I2} → + ∃∃I1,L1. ⦃G, L1⦄ ⊢ ⬈[h, #i] L2 & Y1 = L1.ⓘ{I1}. +/2 width=2 by rex_inv_lref_bind_dx/ qed-. + +lemma rpx_inv_gref_bind_sn: ∀h,I1,G,Y2,L1,l. ⦃G, L1.ⓘ{I1}⦄ ⊢ ⬈[h, §l] Y2 → + ∃∃I2,L2. ⦃G, L1⦄ ⊢ ⬈[h, §l] L2 & Y2 = L2.ⓘ{I2}. +/2 width=2 by rex_inv_gref_bind_sn/ qed-. + +lemma rpx_inv_gref_bind_dx: ∀h,I2,G,Y1,L2,l. ⦃G, Y1⦄ ⊢ ⬈[h, §l] L2.ⓘ{I2} → + ∃∃I1,L1. ⦃G, L1⦄ ⊢ ⬈[h, §l] L2 & Y1 = L1.ⓘ{I1}. +/2 width=2 by rex_inv_gref_bind_dx/ qed-. + +(* Basic forward lemmas *****************************************************) + +lemma rpx_fwd_pair_sn: ∀h,I,G,L1,L2,V,T. + ⦃G, L1⦄ ⊢ ⬈[h, ②{I}V.T] L2 → ⦃G, L1⦄ ⊢ ⬈[h, V] L2. +/2 width=3 by rex_fwd_pair_sn/ qed-. + +lemma rpx_fwd_bind_dx: ∀h,p,I,G,L1,L2,V,T. + ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V.T] L2 → ⦃G, L1.ⓑ{I}V⦄ ⊢ ⬈[h, T] L2.ⓑ{I}V. +/2 width=2 by rex_fwd_bind_dx/ qed-. + +lemma rpx_fwd_flat_dx: ∀h,I,G,L1,L2,V,T. + ⦃G, L1⦄ ⊢ ⬈[h, ⓕ{I}V.T] L2 → ⦃G, L1⦄ ⊢ ⬈[h, T] L2. +/2 width=3 by rex_fwd_flat_dx/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_fqup.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_fqup.ma new file mode 100644 index 000000000..1b49ca7d3 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_fqup.ma @@ -0,0 +1,39 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rex_fqup.ma". +include "basic_2/rt_transition/rpx.ma". + +(* UNBOUND PARALLEL RT-TRANSITION FOR REFERRED LOCAL ENVIRONMENTS ***********) + +(* Advanced properties ******************************************************) + +lemma rpx_refl: ∀h,G,T. reflexive … (rpx h G T). +/2 width=1 by rex_refl/ qed. + +lemma rpx_pair_refl: ∀h,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ⬈[h] V2 → + ∀I,T. ⦃G, L.ⓑ{I}V1⦄ ⊢ ⬈[h, T] L.ⓑ{I}V2. +/2 width=1 by rex_pair_refl/ qed. + +(* Advanced inversion lemmas ************************************************) + +lemma rpx_inv_bind_void: ∀h,p,I,G,L1,L2,V,T. ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V.T] L2 → + ∧∧ ⦃G, L1⦄ ⊢ ⬈[h, V] L2 & ⦃G, L1.ⓧ⦄ ⊢ ⬈[h, T] L2.ⓧ. +/2 width=3 by rex_inv_bind_void/ qed-. + +(* Advanced forward lemmas **************************************************) + +lemma rpx_fwd_bind_dx_void: ∀h,p,I,G,L1,L2,V,T. + ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V.T] L2 → ⦃G, L1.ⓧ⦄ ⊢ ⬈[h, T] L2.ⓧ. +/2 width=4 by rex_fwd_bind_dx_void/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_fsle.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_fsle.ma new file mode 100644 index 000000000..05b8ba41c --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_fsle.ma @@ -0,0 +1,136 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/fsle_drops.ma". +include "basic_2/static/rex_fsle.ma". +include "basic_2/rt_transition/rpx_length.ma". +include "basic_2/rt_transition/rpx_fqup.ma". + +(* UNBOUND PARALLEL RT-TRANSITION FOR REFERRED LOCAL ENVIRONMENTS ***********) + +(* Forward lemmas with free variables inclusion for restricted closures *****) + +(* Note: "⦃L2, T1⦄ ⊆ ⦃L2, T0⦄" does not hold *) +(* Note: Take L0 = K0.ⓓ(ⓝW.V), L2 = K0.ⓓW, T0 = #0, T1 = ⬆*[1]V *) +(* Note: This invalidates rpxs_cpx_conf: "∀h,G. s_r_confluent1 … (cpx h G) (rpxs h G)" *) +lemma rpx_cpx_conf_fsge (h) (G): ∀L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ⬈[h] T1 → + ∀L2. ⦃G, L0⦄ ⊢⬈[h, T0] L2 → ⦃L2, T1⦄ ⊆ ⦃L0, T0⦄. +#h #G0 #L0 #T0 @(fqup_wf_ind_eq (Ⓕ) … G0 L0 T0) -G0 -L0 -T0 +#G #L #T #IH #G0 #L0 * * +[ #s #HG #HL #HT #X #HX #Y #HY destruct -IH + elim (cpx_inv_sort1 … HX) -HX #H destruct + lapply (rpx_fwd_length … HY) -HY #H0 + /2 width=1 by fsle_sort_bi/ +| * [| #i ] #HG #HL #HT #X #HX #Y #HY destruct + [ elim (cpx_inv_zero1 … HX) -HX + [ #H destruct + elim (rpx_inv_zero_length … HY) -HY * + [ #H1 #H2 destruct -IH // + | #I #K0 #K2 #V0 #V2 #HK02 #HV02 #H1 #H2 destruct + lapply (rpx_fwd_length … HK02) #H0 + /4 width=4 by fsle_pair_bi, fqu_fqup, fqu_lref_O/ + | #I #K0 #K2 #HK02 #H1 #H2 destruct -IH + /2 width=1 by fsle_unit_bi/ + ] + | * #I0 #K0 #V0 #V1 #HV01 #HV1X #H destruct + elim (rpx_inv_zero_pair_sn … HY) -HY #K2 #V2 #HK02 #HV02 #H destruct + lapply (rpx_fwd_length … HK02) #H0 + /4 width=4 by fsle_lifts_SO_sn, fqu_fqup, fqu_lref_O/ + ] + | elim (cpx_inv_lref1 … HX) -HX + [ #H destruct + elim (rpx_inv_lref … HY) -HY * + [ #H0 #H1 destruct // + | #I0 #I2 #K0 #K2 #HK02 #H1 #H2 destruct + lapply (rpx_fwd_length … HK02) #H0 + /4 width=5 by fsle_lifts_SO, fqu_fqup/ + ] + | * #I0 #K0 #V1 #HV1 #HV1X #H0 destruct + elim (rpx_inv_lref_bind_sn … HY) -HY #I2 #K2 #HK02 #H destruct + lapply (rpx_fwd_length … HK02) #H0 + /4 width=5 by fsle_lifts_SO, fqu_fqup/ + ] + ] +| #l #HG #HL #HT #X #HX #Y #HY destruct -IH + >(cpx_inv_gref1 … HX) -X + lapply (rpx_fwd_length … HY) -HY #H0 + /2 width=1 by fsle_gref_bi/ +| #p #I #V0 #T0 #HG #HL #HT #X #HX #Y #HY destruct + elim (rpx_inv_bind … V0 ? HY) -HY #HV0 #HT0 + elim (cpx_inv_bind1 … HX) -HX * + [ #V1 #T1 #HV01 #HT01 #H destruct + lapply (rpx_fwd_length … HV0) #H0 + /4 width=6 by fsle_bind_eq, fsle_fwd_pair_sn/ + | #T #HT #HXT #H1 #H2 destruct + lapply (rpx_fwd_length … HV0) #H0 + /3 width=8 by fsle_inv_lifts_sn/ + ] +| #I #V0 #X0 #HG #HL #HT #X #HX #Y #HY destruct + elim (rex_inv_flat … HY) -HY #HV0 #HX0 + elim (cpx_inv_flat1 … HX) -HX * + [ #V1 #T1 #HV01 #HT01 #H destruct + /3 width=4 by fsle_flat/ + | #HX #H destruct + /4 width=4 by fsle_flat_dx_dx/ + | #HX #H destruct + /4 width=4 by fsle_flat_dx_sn/ + | #p #V1 #W0 #W1 #T0 #T1 #HV01 #HW01 #HT01 #H1 #H2 #H3 destruct + elim (rpx_inv_bind … W0 ? HX0) -HX0 #HW0 #HT0 + lapply (rpx_fwd_length … HV0) #H0 + lapply (IH … HV01 … HV0) -HV01 -HV0 // #HV + lapply (IH … HW01 … HW0) -HW01 -HW0 // #HW + lapply (IH … HT01 … HT0) -HT01 -HT0 -IH // #HT + lapply (fsle_fwd_pair_sn … HT) -HT #HT + @fsle_bind_sn_ge // + [ /4 width=1 by fsle_flat_sn, fsle_flat_dx_dx, fsle_flat_dx_sn, fsle_bind_dx_sn/ + | /3 width=1 by fsle_flat_dx_dx, fsle_shift/ + ] + | #p #V1 #X1 #W0 #W1 #T0 #T1 #HV01 #HVX1 #HW01 #HT01 #H1 #H2 #H3 destruct + elim (rpx_inv_bind … W0 ? HX0) -HX0 #HW0 #HT0 + lapply (rpx_fwd_length … HV0) #H0 + lapply (IH … HV01 … HV0) -HV01 -HV0 // #HV + lapply (IH … HW01 … HW0) -HW01 -HW0 // #HW + lapply (IH … HT01 … HT0) -HT01 -HT0 -IH // #HT + lapply (fsle_fwd_pair_sn … HT) -HT #HT + @fsle_bind_sn_ge // + [ /3 width=1 by fsle_flat_dx_dx, fsle_bind_dx_sn/ + | /4 width=3 by fsle_flat_sn, fsle_flat_dx_sn, fsle_flat_dx_dx, fsle_shift, fsle_lifts_sn/ + ] + ] +] +qed-. + +lemma rpx_fsge_comp (h) (G): rex_fsge_compatible (cpx h G). +/2 width=4 by rpx_cpx_conf_fsge/ qed-. + +(**) (* this section concerns cpx *) +(* Properties with generic extension on referred entries ********************) + +(* Basic_2A1: uses: cpx_frees_trans *) +lemma cpx_fsge_comp (h) (G): R_fsge_compatible (cpx h G). +/2 width=4 by rpx_cpx_conf_fsge/ qed-. + +(* Note: lemma 1000 *) +(* Basic_2A1: uses: cpx_llpx_sn_conf *) +lemma cpx_rex_conf (R) (h) (G): s_r_confluent1 … (cpx h G) (rex R). +/3 width=3 by fsge_rex_trans, cpx_fsge_comp/ qed-. + +(* Advanced properties ******************************************************) + +lemma rpx_cpx_conf (h) (G): s_r_confluent1 … (cpx h G) (rpx h G). +/2 width=5 by cpx_rex_conf/ qed-. + +lemma rpx_cpx_conf_fsge_dx (h) (G): ∀L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ⬈[h] T1 → + ∀L2. ⦃G, L0⦄ ⊢⬈[h, T0] L2 → ⦃L2, T1⦄ ⊆ ⦃L0, T1⦄. +/3 width=5 by rpx_cpx_conf, rpx_fsge_comp/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_length.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_length.ma new file mode 100644 index 000000000..fe095d1df --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_length.ma @@ -0,0 +1,33 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rex_length.ma". +include "basic_2/rt_transition/rpx.ma". + +(* UNBOUND PARALLEL RT-TRANSITION FOR REFERRED LOCAL ENVIRONMENTS ***********) + +(* Forward lemmas with length for local environments ************************) + +lemma rpx_fwd_length: ∀h,G,L1,L2,T. ⦃G, L1⦄ ⊢ ⬈[h, T] L2 → |L1| = |L2|. +/2 width=3 by rex_fwd_length/ qed-. + +(* Inversion lemmas with length for local environments **********************) + +lemma rpx_inv_zero_length: ∀h,G,Y1,Y2. ⦃G, Y1⦄ ⊢ ⬈[h, #0] Y2 → + ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆ + | ∃∃I,L1,L2,V1,V2. ⦃G, L1⦄ ⊢ ⬈[h, V1] L2 & + ⦃G, L1⦄ ⊢ V1 ⬈[h] V2 & + Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2 + | ∃∃I,L1,L2. |L1| = |L2| & Y1 = L1.ⓤ{I} & Y2 = L2.ⓤ{I}. +/2 width=1 by rex_inv_zero_length/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_lpx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_lpx.ma new file mode 100644 index 000000000..585078496 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_lpx.ma @@ -0,0 +1,35 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rex_lex.ma". +include "basic_2/rt_transition/rpx_fsle.ma". +include "basic_2/rt_transition/lpx.ma". + +(* UNBOUND PARALLEL RT-TRANSITION FOR REFERRED LOCAL ENVIRONMENTS ***********) + +(* Properties with syntactic equivalence for referred local environments ****) + +lemma fleq_rpx (h) (G): ∀L1,L2,T. L1 ≡[T] L2 → ⦃G, L1⦄ ⊢ ⬈[h, T] L2. +/2 width=1 by req_fwd_rex/ qed. + +(* Properties with unbound parallel rt-transition for full local envs *******) + +lemma lpx_rpx: ∀h,G,L1,L2,T. ⦃G, L1⦄ ⊢ ⬈[h] L2 → ⦃G, L1⦄ ⊢ ⬈[h, T] L2. +/2 width=1 by rex_lex/ qed. + +(* Inversion lemmas with unbound parallel rt-transition for full local envs *) + +lemma rpx_inv_lpx_req: ∀h,G,L1,L2,T. ⦃G, L1⦄ ⊢ ⬈[h, T] L2 → + ∃∃L. ⦃G, L1⦄ ⊢ ⬈[h] L & L ≡[T] L2. +/3 width=3 by rpx_fsge_comp, rex_inv_lex_req/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_rdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_rdeq.ma new file mode 100644 index 000000000..e2513c96f --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_rdeq.ma @@ -0,0 +1,168 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rdeq_fqup.ma". +include "basic_2/static/rdeq_rdeq.ma". +include "basic_2/rt_transition/rpx_fsle.ma". + +(* UNBOUND PARALLEL RT-TRANSITION FOR REFERRED LOCAL ENVIRONMENTS ***********) + +(* Properties with degree-based equivalence for local environments **********) + +lemma rpx_pair_sn_split: ∀h,G,L1,L2,V. ⦃G, L1⦄ ⊢ ⬈[h, V] L2 → ∀o,I,T. + ∃∃L. ⦃G, L1⦄ ⊢ ⬈[h, ②{I}V.T] L & L ≛[h, o, V] L2. +/3 width=5 by rpx_fsge_comp, rex_pair_sn_split/ qed-. + +lemma rpx_flat_dx_split: ∀h,G,L1,L2,T. ⦃G, L1⦄ ⊢ ⬈[h, T] L2 → ∀o,I,V. + ∃∃L. ⦃G, L1⦄ ⊢ ⬈[h, ⓕ{I}V.T] L & L ≛[h, o, T] L2. +/3 width=5 by rpx_fsge_comp, rex_flat_dx_split/ qed-. + +lemma rpx_bind_dx_split: ∀h,I,G,L1,L2,V1,T. ⦃G, L1.ⓑ{I}V1⦄ ⊢ ⬈[h, T] L2 → ∀o,p. + ∃∃L,V. ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V1.T] L & L.ⓑ{I}V ≛[h, o, T] L2 & ⦃G, L1⦄ ⊢ V1 ⬈[h] V. +/3 width=5 by rpx_fsge_comp, rex_bind_dx_split/ qed-. + +lemma rpx_bind_dx_split_void: ∀h,G,K1,L2,T. ⦃G, K1.ⓧ⦄ ⊢ ⬈[h, T] L2 → ∀o,p,I,V. + ∃∃K2. ⦃G, K1⦄ ⊢ ⬈[h, ⓑ{p,I}V.T] K2 & K2.ⓧ ≛[h, o, T] L2. +/3 width=5 by rpx_fsge_comp, rex_bind_dx_split_void/ qed-. + +lemma rpx_tdeq_conf: ∀h,o,G. s_r_confluent1 … (cdeq h o) (rpx h G). +/2 width=5 by tdeq_rex_conf/ qed-. + +lemma rpx_tdeq_div: ∀h,o,T1,T2. T1 ≛[h, o] T2 → + ∀G,L1,L2. ⦃G, L1⦄ ⊢ ⬈[h, T2] L2 → ⦃G, L1⦄ ⊢ ⬈[h, T1] L2. +/2 width=5 by tdeq_rex_div/ qed-. + +lemma cpx_tdeq_conf_sex: ∀h,o,G. R_confluent2_rex … (cpx h G) (cdeq h o) (cpx h G) (cdeq h o). +#h #o #G #L0 #T0 #T1 #H @(cpx_ind … H) -G -L0 -T0 -T1 /2 width=3 by ex2_intro/ +[ #G #L0 #s0 #X0 #H0 #L1 #HL01 #L2 #HL02 + elim (tdeq_inv_sort1 … H0) -H0 #s1 #d1 #Hs0 #Hs1 #H destruct + /4 width=3 by tdeq_sort, deg_next, ex2_intro/ +| #I #G #K0 #V0 #V1 #W1 #_ #IH #HVW1 #T2 #H0 #L1 #H1 #L2 #H2 + >(tdeq_inv_lref1 … H0) -H0 + elim (rpx_inv_zero_pair_sn … H1) -H1 #K1 #X1 #HK01 #HX1 #H destruct + elim (rdeq_inv_zero_pair_sn … H2) -H2 #K2 #X2 #HK02 #HX2 #H destruct + elim (IH X2 … HK01 … HK02) // -K0 -V0 #V #HV1 #HV2 + elim (tdeq_lifts_sn … HV1 … HVW1) -V1 /3 width=5 by cpx_delta, ex2_intro/ +| #I0 #G #K0 #V1 #W1 #i #_ #IH #HVW1 #T2 #H0 #L1 #H1 #L2 #H2 + >(tdeq_inv_lref1 … H0) -H0 + elim (rpx_inv_lref_bind_sn … H1) -H1 #I1 #K1 #HK01 #H destruct + elim (rdeq_inv_lref_bind_sn … H2) -H2 #I2 #K2 #HK02 #H destruct + elim (IH … HK01 … HK02) [|*: //] -K0 #V #HV1 #HV2 + elim (tdeq_lifts_sn … HV1 … HVW1) -V1 /3 width=5 by cpx_lref, ex2_intro/ +| #p #I #G #L0 #V0 #V1 #T0 #T1 #_ #_ #IHV #IHT #X0 #H0 #L1 #H1 #L2 #H2 + elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #HV02 #HT02 #H destruct + elim (rpx_inv_bind … H1) -H1 #HL01 #H1 + elim (rdeq_inv_bind … H2) -H2 #HL02 #H2 + lapply (rdeq_bind_repl_dx … H2 (BPair I V2) ?) -H2 /2 width=1 by ext2_pair/ #H2 + elim (IHV … HV02 … HL01 … HL02) -IHV -HV02 -HL01 -HL02 + elim (IHT … HT02 … H1 … H2) -L0 -T0 + /3 width=5 by cpx_bind, tdeq_pair, ex2_intro/ +| #I #G #L0 #V0 #V1 #T0 #T1 #_ #_ #IHV #IHT #X0 #H0 #L1 #H1 #L2 #H2 + elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #HV02 #HT02 #H destruct + elim (rpx_inv_flat … H1) -H1 #HL01 #H1 + elim (rdeq_inv_flat … H2) -H2 #HL02 #H2 + elim (IHV … HV02 … HL01 … HL02) -IHV -HV02 -HL01 -HL02 + elim (IHT … HT02 … H1 … H2) -L0 -V0 -T0 + /3 width=5 by cpx_flat, tdeq_pair, ex2_intro/ +| #G #L0 #V0 #T0 #T1 #U1 #_ #IH #HUT1 #X0 #H0 #L1 #H1 #L2 #H2 + elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #HV02 #HT02 #H destruct + elim (rpx_inv_bind … H1) -H1 #HL01 #H1 + elim (rdeq_inv_bind … H2) -H2 #HL02 #H2 + lapply (rdeq_bind_repl_dx … H2 (BPair Abbr V2) ?) -H2 /2 width=1 by ext2_pair/ -HV02 #H2 + elim (IH … HT02 … H1 … H2) -L0 -T0 #T #HT1 + elim (tdeq_inv_lifts_sn … HT1 … HUT1) -T1 + /3 width=5 by cpx_zeta, ex2_intro/ +| #G #L0 #V0 #T0 #T1 #_ #IH #X0 #H0 #L1 #H1 #L2 #H2 + elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #_ #HT02 #H destruct + elim (rpx_inv_flat … H1) -H1 #HL01 #H1 + elim (rdeq_inv_flat … H2) -H2 #HL02 #H2 + elim (IH … HT02 … H1 … H2) -L0 -V0 -T0 + /3 width=3 by cpx_eps, ex2_intro/ +| #G #L0 #V0 #T0 #T1 #_ #IH #X0 #H0 #L1 #H1 #L2 #H2 + elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #HV02 #_ #H destruct + elim (rpx_inv_flat … H1) -H1 #HL01 #H1 + elim (rdeq_inv_flat … H2) -H2 #HL02 #H2 + elim (IH … HV02 … HL01 … HL02) -L0 -V0 -T1 + /3 width=3 by cpx_ee, ex2_intro/ +| #p #G #L0 #V0 #V1 #W0 #W1 #T0 #T1 #_ #_ #_ #IHV #IHW #IHT #X0 #H0 #L1 #H1 #L2 #H2 + elim (tdeq_inv_pair1 … H0) -H0 #V2 #X #HV02 #H0 #H destruct + elim (tdeq_inv_pair1 … H0) -H0 #W2 #T2 #HW02 #HT02 #H destruct + elim (rpx_inv_flat … H1) -H1 #H1LV0 #H1 + elim (rpx_inv_bind … H1) -H1 #H1LW0 #H1LT0 + elim (rdeq_inv_flat … H2) -H2 #H2LV0 #H2 + elim (rdeq_inv_bind … H2) -H2 #H2LW0 #H2LT0 + lapply (rdeq_bind_repl_dx … H2LT0 (BPair Abst W2) ?) -H2LT0 /2 width=1 by ext2_pair/ #H2LT0 + elim (IHV … HV02 … H1LV0 … H2LV0) -IHV -HV02 -H1LV0 -H2LV0 + elim (IHW … HW02 … H1LW0 … H2LW0) -IHW -HW02 -H1LW0 -H2LW0 + elim (IHT … HT02 … H1LT0 … H2LT0) -L0 -V0 -T0 + /4 width=7 by cpx_beta, tdeq_pair, ex2_intro/ (* note: 2 tdeq_pair *) +| #p #G #L0 #V0 #V1 #U1 #W0 #W1 #T0 #T1 #_ #_ #_ #IHV #IHW #IHT #HVU1 #X0 #H0 #L1 #H1 #L2 #H2 + elim (tdeq_inv_pair1 … H0) -H0 #V2 #X #HV02 #H0 #H destruct + elim (tdeq_inv_pair1 … H0) -H0 #W2 #T2 #HW02 #HT02 #H destruct + elim (rpx_inv_flat … H1) -H1 #H1LV0 #H1 + elim (rpx_inv_bind … H1) -H1 #H1LW0 #H1LT0 + elim (rdeq_inv_flat … H2) -H2 #H2LV0 #H2 + elim (rdeq_inv_bind … H2) -H2 #H2LW0 #H2LT0 + lapply (rdeq_bind_repl_dx … H2LT0 (BPair Abbr W2) ?) -H2LT0 /2 width=1 by ext2_pair/ #H2LT0 + elim (IHV … HV02 … H1LV0 … H2LV0) -IHV -HV02 -H1LV0 -H2LV0 #V #HV1 + elim (IHW … HW02 … H1LW0 … H2LW0) -IHW -HW02 -H1LW0 -H2LW0 + elim (IHT … HT02 … H1LT0 … H2LT0) -L0 -V0 -T0 + elim (tdeq_lifts_sn … HV1 … HVU1) -V1 + /4 width=9 by cpx_theta, tdeq_pair, ex2_intro/ (* note: 2 tdeq_pair *) +] +qed-. + +lemma cpx_tdeq_conf: ∀h,o,G,L. ∀T0:term. ∀T1. ⦃G, L⦄ ⊢ T0 ⬈[h] T1 → + ∀T2. T0 ≛[h, o] T2 → + ∃∃T. T1 ≛[h, o] T & ⦃G, L⦄ ⊢ T2 ⬈[h] T. +#h #o #G #L #T0 #T1 #HT01 #T2 #HT02 +elim (cpx_tdeq_conf_sex … HT01 … HT02 L … L) -HT01 -HT02 +/2 width=3 by rex_refl, ex2_intro/ +qed-. + +lemma tdeq_cpx_trans: ∀h,o,G,L,T2. ∀T0:term. T2 ≛[h, o] T0 → + ∀T1. ⦃G, L⦄ ⊢ T0 ⬈[h] T1 → + ∃∃T. ⦃G, L⦄ ⊢ T2 ⬈[h] T & T ≛[h, o] T1. +#h #o #G #L #T2 #T0 #HT20 #T1 #HT01 +elim (cpx_tdeq_conf … HT01 T2) -HT01 /3 width=3 by tdeq_sym, ex2_intro/ +qed-. + +(* Basic_2A1: uses: cpx_lleq_conf *) +lemma cpx_rdeq_conf: ∀h,o,G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ⬈[h] T1 → + ∀L2. L0 ≛[h, o, T0] L2 → + ∃∃T. ⦃G, L2⦄ ⊢ T0 ⬈[h] T & T1 ≛[h, o] T. +#h #o #G #L0 #T0 #T1 #HT01 #L2 #HL02 +elim (cpx_tdeq_conf_sex … HT01 T0 … L0 … HL02) -HT01 -HL02 +/2 width=3 by rex_refl, ex2_intro/ +qed-. + +(* Basic_2A1: uses: lleq_cpx_trans *) +lemma rdeq_cpx_trans: ∀h,o,G,L2,L0,T0. L2 ≛[h, o, T0] L0 → + ∀T1. ⦃G, L0⦄ ⊢ T0 ⬈[h] T1 → + ∃∃T. ⦃G, L2⦄ ⊢ T0 ⬈[h] T & T ≛[h, o] T1. +#h #o #G #L2 #L0 #T0 #HL20 #T1 #HT01 +elim (cpx_rdeq_conf … o … HT01 L2) -HT01 +/3 width=3 by rdeq_sym, tdeq_sym, ex2_intro/ +qed-. + +lemma rpx_rdeq_conf: ∀h,o,G,T. confluent2 … (rpx h G T) (rdeq h o T). +/3 width=6 by rpx_fsge_comp, rdeq_fsge_comp, cpx_tdeq_conf_sex, rex_conf/ qed-. + +lemma rdeq_rpx_trans: ∀h,o,G,T,L2,K2. ⦃G, L2⦄ ⊢ ⬈[h, T] K2 → + ∀L1. L1 ≛[h, o, T] L2 → + ∃∃K1. ⦃G, L1⦄ ⊢ ⬈[h, T] K1 & K1 ≛[h, o, T] K2. +#h #o #G #T #L2 #K2 #HLK2 #L1 #HL12 +elim (rpx_rdeq_conf … o … HLK2 L1) +/3 width=3 by rdeq_sym, ex2_intro/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_rpx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_rpx.ma new file mode 100644 index 000000000..b5ccfa88c --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/rpx_rpx.ma @@ -0,0 +1,34 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rex_rex.ma". +include "basic_2/rt_transition/rpx.ma". + +(* UNBOUND PARALLEL RT-TRANSITION FOR REFERRED LOCAL ENVIRONMENTS ***********) + +(* Main properties **********************************************************) + +theorem rpx_bind: ∀h,G,L1,L2,V1. ⦃G, L1⦄ ⊢ ⬈[h, V1] L2 → + ∀I,V2,T. ⦃G, L1.ⓑ{I}V1⦄ ⊢ ⬈[h, T] L2.ⓑ{I}V2 → + ∀p. ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V1.T] L2. +/2 width=2 by rex_bind/ qed. + +theorem rpx_flat: ∀h,G,L1,L2,V. ⦃G, L1⦄ ⊢ ⬈[h, V] L2 → + ∀I,T. ⦃G, L1⦄ ⊢ ⬈[h, T] L2 → ⦃G, L1⦄ ⊢ ⬈[h, ⓕ{I}V.T] L2. +/2 width=1 by rex_flat/ qed. + +theorem rpx_bind_void: ∀h,G,L1,L2,V. ⦃G, L1⦄ ⊢ ⬈[h, V] L2 → + ∀T. ⦃G, L1.ⓧ⦄ ⊢ ⬈[h, T] L2.ⓧ → + ∀p,I. ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V.T] L2. +/2 width=1 by rex_bind_void/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_fdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_fdeq.ma new file mode 100644 index 000000000..13b5acbd6 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_fdeq.ma @@ -0,0 +1,25 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/fdeq.ma". +include "basic_2/static/aaa_rdeq.ma". + +(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) + +(* Properties with degree-based equivalence on referred entries *************) + +lemma aaa_fdeq_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → + ∀A. ⦃G1, L1⦄ ⊢ T1 ⁝ A → ⦃G2, L2⦄ ⊢ T2 ⁝ A. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 +/2 width=7 by aaa_tdeq_conf_rdeq/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_ffdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_ffdeq.ma deleted file mode 100644 index f9495e62d..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_ffdeq.ma +++ /dev/null @@ -1,25 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/ffdeq.ma". -include "basic_2/static/aaa_lfdeq.ma". - -(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) - -(* Properties with degree-based equivalence on referred entries *************) - -lemma aaa_ffdeq_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → - ∀A. ⦃G1, L1⦄ ⊢ T1 ⁝ A → ⦃G2, L2⦄ ⊢ T2 ⁝ A. -#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 -/2 width=7 by aaa_tdeq_conf_lfdeq/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lfdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lfdeq.ma deleted file mode 100644 index a01878128..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lfdeq.ma +++ /dev/null @@ -1,45 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfdeq.ma". -include "basic_2/static/aaa.ma". - -(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) - -(* Properties with degree-based equivalence on referred entries *************) - -lemma aaa_tdeq_conf_lfdeq: ∀h,o,G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → ∀T2. T1 ≛[h, o] T2 → - ∀L2. L1 ≛[h, o, T1] L2 → ⦃G, L2⦄ ⊢ T2 ⁝ A. -#h #o #G #L1 #T1 #A #H elim H -G -L1 -T1 -A -[ #G #L1 #s1 #X #H1 elim (tdeq_inv_sort1 … H1) -H1 // -| #I #G #L1 #V1 #B #_ #IH #X #H1 >(tdeq_inv_lref1 … H1) -H1 - #Y #H2 elim (lfdeq_inv_zero_pair_sn … H2) -H2 - #L2 #V2 #HL12 #HV12 #H destruct /3 width=1 by aaa_zero/ -| #I #G #L1 #A #i #_ #IH #X #H1 >(tdeq_inv_lref1 … H1) -H1 - #Y #H2 elim (lfdeq_inv_lref_bind_sn … H2) -H2 - #J #L2 #HL12 #H destruct /3 width=1 by aaa_lref/ -| #p #G #L1 #V1 #T1 #B #A #_ #_ #IHV #IHT #X #H1 elim (tdeq_inv_pair1 … H1) -H1 - #V2 #T2 #HV12 #HT12 #H #L2 #H2 elim (lfdeq_inv_bind … H2) -H2 destruct - /5 width=2 by aaa_abbr, lfdeq_bind_repl_dx, ext2_pair/ -| #p #G #L1 #V1 #T1 #B #A #_ #_ #IHV #IHT #X #H1 elim (tdeq_inv_pair1 … H1) -H1 - #V2 #T2 #HV12 #HT12 #H #L2 #H2 elim (lfdeq_inv_bind … H2) -H2 destruct - /5 width=2 by aaa_abst, lfdeq_bind_repl_dx, ext2_pair/ -| #G #L1 #V1 #T1 #B #A #_ #_ #IHV #IHT #X #H1 elim (tdeq_inv_pair1 … H1) -H1 - #V2 #T2 #HV12 #HT12 #H #L2 #H2 elim (lfdeq_inv_flat … H2) -H2 destruct - /3 width=3 by aaa_appl/ -| #G #L1 #V1 #T1 #A #_ #_ #IHV #IHT #X #H1 elim (tdeq_inv_pair1 … H1) -H1 - #V2 #T2 #HV12 #HT12 #H #L2 #H2 elim (lfdeq_inv_flat … H2) -H2 destruct - /3 width=1 by aaa_cast/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_rdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_rdeq.ma new file mode 100644 index 000000000..d55f97ce7 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_rdeq.ma @@ -0,0 +1,45 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rdeq.ma". +include "basic_2/static/aaa.ma". + +(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) + +(* Properties with degree-based equivalence on referred entries *************) + +lemma aaa_tdeq_conf_rdeq: ∀h,o,G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → ∀T2. T1 ≛[h, o] T2 → + ∀L2. L1 ≛[h, o, T1] L2 → ⦃G, L2⦄ ⊢ T2 ⁝ A. +#h #o #G #L1 #T1 #A #H elim H -G -L1 -T1 -A +[ #G #L1 #s1 #X #H1 elim (tdeq_inv_sort1 … H1) -H1 // +| #I #G #L1 #V1 #B #_ #IH #X #H1 >(tdeq_inv_lref1 … H1) -H1 + #Y #H2 elim (rdeq_inv_zero_pair_sn … H2) -H2 + #L2 #V2 #HL12 #HV12 #H destruct /3 width=1 by aaa_zero/ +| #I #G #L1 #A #i #_ #IH #X #H1 >(tdeq_inv_lref1 … H1) -H1 + #Y #H2 elim (rdeq_inv_lref_bind_sn … H2) -H2 + #J #L2 #HL12 #H destruct /3 width=1 by aaa_lref/ +| #p #G #L1 #V1 #T1 #B #A #_ #_ #IHV #IHT #X #H1 elim (tdeq_inv_pair1 … H1) -H1 + #V2 #T2 #HV12 #HT12 #H #L2 #H2 elim (rdeq_inv_bind … H2) -H2 destruct + /5 width=2 by aaa_abbr, rdeq_bind_repl_dx, ext2_pair/ +| #p #G #L1 #V1 #T1 #B #A #_ #_ #IHV #IHT #X #H1 elim (tdeq_inv_pair1 … H1) -H1 + #V2 #T2 #HV12 #HT12 #H #L2 #H2 elim (rdeq_inv_bind … H2) -H2 destruct + /5 width=2 by aaa_abst, rdeq_bind_repl_dx, ext2_pair/ +| #G #L1 #V1 #T1 #B #A #_ #_ #IHV #IHT #X #H1 elim (tdeq_inv_pair1 … H1) -H1 + #V2 #T2 #HV12 #HT12 #H #L2 #H2 elim (rdeq_inv_flat … H2) -H2 destruct + /3 width=3 by aaa_appl/ +| #G #L1 #V1 #T1 #A #_ #_ #IHV #IHT #X #H1 elim (tdeq_inv_pair1 … H1) -H1 + #V2 #T2 #HV12 #HT12 #H #L2 #H2 elim (rdeq_inv_flat … H2) -H2 destruct + /3 width=1 by aaa_cast/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/fdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/static/fdeq.ma new file mode 100644 index 000000000..6d8dd1589 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/fdeq.ma @@ -0,0 +1,52 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/notation/relations/stareqsn_8.ma". +include "basic_2/syntax/genv.ma". +include "basic_2/static/rdeq.ma". + +(* DEGREE-BASED EQUIVALENCE FOR CLOSURES ON REFERRED ENTRIES ****************) + +inductive fdeq (h) (o) (G) (L1) (T1): relation3 genv lenv term ≝ +| fdeq_intro_sn: ∀L2,T2. L1 ≛[h, o, T1] L2 → T1 ≛[h, o] T2 → + fdeq h o G L1 T1 G L2 T2 +. + +interpretation + "degree-based equivalence on referred entries (closure)" + 'StarEqSn h o G1 L1 T1 G2 L2 T2 = (fdeq h o G1 L1 T1 G2 L2 T2). + +(* Basic_properties *********************************************************) + +lemma fdeq_intro_dx (h) (o) (G): ∀L1,L2,T2. L1 ≛[h, o, T2] L2 → + ∀T1. T1 ≛[h, o] T2 → ⦃G, L1, T1⦄ ≛[h, o] ⦃G, L2, T2⦄. +/3 width=3 by fdeq_intro_sn, tdeq_rdeq_div/ qed. + +(* Basic inversion lemmas ***************************************************) + +lemma fdeq_inv_gen_sn: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → + ∧∧ G1 = G2 & L1 ≛[h, o, T1] L2 & T1 ≛[h, o] T2. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 /2 width=1 by and3_intro/ +qed-. + +lemma fdeq_inv_gen_dx: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → + ∧∧ G1 = G2 & L1 ≛[h, o, T2] L2 & T1 ≛[h, o] T2. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 +/3 width=3 by tdeq_rdeq_conf, and3_intro/ +qed-. + +(* Basic_2A1: removed theorems 6: + fleq_refl fleq_sym fleq_inv_gen + fleq_trans fleq_canc_sn fleq_canc_dx +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/fdeq_fdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/static/fdeq_fdeq.ma new file mode 100644 index 000000000..13ab8ff2f --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/fdeq_fdeq.ma @@ -0,0 +1,52 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rdeq_rdeq.ma". +include "basic_2/static/fdeq.ma". + +(* DEGREE-BASED EQUIVALENCE FOR CLOSURES ON REFERRED ENTRIES ****************) + +(* Advanced properties ******************************************************) + +lemma fdeq_sym: ∀h,o. tri_symmetric … (fdeq h o). +#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G1 -L1 -T1 +/3 width=1 by fdeq_intro_dx, rdeq_sym, tdeq_sym/ +qed-. + +(* Main properties **********************************************************) + +theorem fdeq_trans: ∀h,o. tri_transitive … (fdeq h o). +#h #o #G1 #G #L1 #L #T1 #T * -G -L -T +#L #T #HL1 #HT1 #G2 #L2 #T2 * -G2 -L2 -T2 +/4 width=5 by fdeq_intro_sn, rdeq_trans, tdeq_rdeq_div, tdeq_trans/ +qed-. + +theorem fdeq_canc_sn: ∀h,o,G,G1,G2,L,L1,L2,T,T1,T2. + ⦃G, L, T⦄ ≛[h, o] ⦃G1, L1, T1⦄→ ⦃G, L, T⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄. +/3 width=5 by fdeq_trans, fdeq_sym/ qed-. + +theorem fdeq_canc_dx: ∀h,o,G1,G2,G,L1,L2,L,T1,T2,T. + ⦃G1, L1, T1⦄ ≛[h, o] ⦃G, L, T⦄ → ⦃G2, L2, T2⦄ ≛[h, o] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄. +/3 width=5 by fdeq_trans, fdeq_sym/ qed-. + +(* Main inversion lemmas with degree-based equivalence on terms *************) + +theorem fdeq_tdneq_repl_dx: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → + ∀U1,U2. ⦃G1, L1, U1⦄ ≛[h, o] ⦃G2, L2, U2⦄ → + (T2 ≛[h, o] U2 → ⊥) → (T1 ≛[h, o] U1 → ⊥). +#h #o #G1 #G2 #L1 #L2 #T1 #T2 #HT #U1 #U2 #HU #HnTU2 #HTU1 +elim (fdeq_inv_gen_sn … HT) -HT #_ #_ #HT +elim (fdeq_inv_gen_sn … HU) -HU #_ #_ #HU +/3 width=5 by tdeq_repl/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/fdeq_fqup.ma b/matita/matita/contribs/lambdadelta/basic_2/static/fdeq_fqup.ma new file mode 100644 index 000000000..e9cc92a39 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/fdeq_fqup.ma @@ -0,0 +1,29 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rdeq_fqup.ma". +include "basic_2/static/fdeq.ma". + +(* DEGREE-BASED EQUIVALENCE FOR CLOSURES ON REFERRED ENTRIES ****************) + +(* Properties with degree-based equivalence for terms ***********************) + +lemma tdeq_fdeq: ∀h,o,T1,T2. T1 ≛[h, o] T2 → + ∀G,L. ⦃G, L, T1⦄ ≛[h, o] ⦃G, L, T2⦄. +/2 width=1 by fdeq_intro_sn/ qed. + +(* Advanced properties ******************************************************) + +lemma fdeq_refl: ∀h,o. tri_reflexive … (fdeq h o). +/2 width=1 by fdeq_intro_sn/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/fdeq_fqus.ma b/matita/matita/contribs/lambdadelta/basic_2/static/fdeq_fqus.ma new file mode 100644 index 000000000..6c2246cad --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/fdeq_fqus.ma @@ -0,0 +1,31 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rdeq_fqus.ma". +include "basic_2/static/fdeq.ma". + +(* DEGREE-BASED EQUIVALENCE FOR CLOSURES ON REFERRED ENTRIES ****************) + +(* Properties with star-iterated structural successor for closures **********) + +lemma fdeq_fqus_trans: ∀h,o,b,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G, L, T⦄ → + ∀G2,L2,T2. ⦃G, L, T⦄ ⊐*[b] ⦃G2, L2, T2⦄ → + ∃∃G,L0,T0. ⦃G1, L1, T1⦄ ⊐*[b] ⦃G, L0, T0⦄ & ⦃G, L0, T0⦄ ≛[h, o] ⦃G2, L2, T2⦄. +#h #o #b #G1 #G #L1 #L #T1 #T #H1 #G2 #L2 #T2 #H2 +elim(fdeq_inv_gen_dx … H1) -H1 #HG #HL1 #HT1 destruct +elim (rdeq_fqus_trans … H2 … HL1) -L #L #T0 #H2 #HT02 #HL2 +elim (tdeq_fqus_trans … H2 … HT1) -T #L0 #T #H2 #HT0 #HL0 +lapply (tdeq_rdeq_conf … HT02 … HL0) -HL0 #HL0 +/4 width=7 by fdeq_intro_dx, rdeq_trans, tdeq_trans, ex2_3_intro/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/fdeq_req.ma b/matita/matita/contribs/lambdadelta/basic_2/static/fdeq_req.ma new file mode 100644 index 000000000..7f98713a0 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/fdeq_req.ma @@ -0,0 +1,27 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rdeq_req.ma". +include "basic_2/static/fdeq.ma". + +(* DEGREE-BASED EQUIVALENCE FOR CLOSURES ON REFERRED ENTRIES ****************) + +(* Properties with syntactic equivalence on referred entries ****************) + +lemma req_rdeq_trans: ∀h,o,L1,L,T1. L1 ≡[T1] L → + ∀G1,G2,L2,T2. ⦃G1, L, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄. +#h #o #L1 #L #T1 #HL1 #G1 #G2 #L2 #T2 #H +elim (fdeq_inv_gen_sn … H) -H #H #HL2 #T12 destruct +/3 width=3 by fdeq_intro_sn, req_rdeq_trans/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/ffdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/static/ffdeq.ma deleted file mode 100644 index ac4a6b169..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/ffdeq.ma +++ /dev/null @@ -1,52 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/notation/relations/stareqsn_8.ma". -include "basic_2/syntax/genv.ma". -include "basic_2/static/lfdeq.ma". - -(* DEGREE-BASED EQUIVALENCE FOR CLOSURES ON REFERRED ENTRIES ****************) - -inductive ffdeq (h) (o) (G) (L1) (T1): relation3 genv lenv term ≝ -| ffdeq_intro_sn: ∀L2,T2. L1 ≛[h, o, T1] L2 → T1 ≛[h, o] T2 → - ffdeq h o G L1 T1 G L2 T2 -. - -interpretation - "degree-based equivalence on referred entries (closure)" - 'StarEqSn h o G1 L1 T1 G2 L2 T2 = (ffdeq h o G1 L1 T1 G2 L2 T2). - -(* Basic_properties *********************************************************) - -lemma ffdeq_intro_dx (h) (o) (G): ∀L1,L2,T2. L1 ≛[h, o, T2] L2 → - ∀T1. T1 ≛[h, o] T2 → ⦃G, L1, T1⦄ ≛[h, o] ⦃G, L2, T2⦄. -/3 width=3 by ffdeq_intro_sn, tdeq_lfdeq_div/ qed. - -(* Basic inversion lemmas ***************************************************) - -lemma ffdeq_inv_gen_sn: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → - ∧∧ G1 = G2 & L1 ≛[h, o, T1] L2 & T1 ≛[h, o] T2. -#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 /2 width=1 by and3_intro/ -qed-. - -lemma ffdeq_inv_gen_dx: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → - ∧∧ G1 = G2 & L1 ≛[h, o, T2] L2 & T1 ≛[h, o] T2. -#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 -/3 width=3 by tdeq_lfdeq_conf, and3_intro/ -qed-. - -(* Basic_2A1: removed theorems 6: - fleq_refl fleq_sym fleq_inv_gen - fleq_trans fleq_canc_sn fleq_canc_dx -*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/ffdeq_ffdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/static/ffdeq_ffdeq.ma deleted file mode 100644 index cbd1284ce..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/ffdeq_ffdeq.ma +++ /dev/null @@ -1,52 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfdeq_lfdeq.ma". -include "basic_2/static/ffdeq.ma". - -(* DEGREE-BASED EQUIVALENCE FOR CLOSURES ON REFERRED ENTRIES ****************) - -(* Advanced properties ******************************************************) - -lemma ffdeq_sym: ∀h,o. tri_symmetric … (ffdeq h o). -#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G1 -L1 -T1 -/3 width=1 by ffdeq_intro_dx, lfdeq_sym, tdeq_sym/ -qed-. - -(* Main properties **********************************************************) - -theorem ffdeq_trans: ∀h,o. tri_transitive … (ffdeq h o). -#h #o #G1 #G #L1 #L #T1 #T * -G -L -T -#L #T #HL1 #HT1 #G2 #L2 #T2 * -G2 -L2 -T2 -/4 width=5 by ffdeq_intro_sn, lfdeq_trans, tdeq_lfdeq_div, tdeq_trans/ -qed-. - -theorem ffdeq_canc_sn: ∀h,o,G,G1,G2,L,L1,L2,T,T1,T2. - ⦃G, L, T⦄ ≛[h, o] ⦃G1, L1, T1⦄→ ⦃G, L, T⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄. -/3 width=5 by ffdeq_trans, ffdeq_sym/ qed-. - -theorem ffdeq_canc_dx: ∀h,o,G1,G2,G,L1,L2,L,T1,T2,T. - ⦃G1, L1, T1⦄ ≛[h, o] ⦃G, L, T⦄ → ⦃G2, L2, T2⦄ ≛[h, o] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄. -/3 width=5 by ffdeq_trans, ffdeq_sym/ qed-. - -(* Main inversion lemmas with degree-based equivalence on terms *************) - -theorem ffdeq_tdneq_repl_dx: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → - ∀U1,U2. ⦃G1, L1, U1⦄ ≛[h, o] ⦃G2, L2, U2⦄ → - (T2 ≛[h, o] U2 → ⊥) → (T1 ≛[h, o] U1 → ⊥). -#h #o #G1 #G2 #L1 #L2 #T1 #T2 #HT #U1 #U2 #HU #HnTU2 #HTU1 -elim (ffdeq_inv_gen_sn … HT) -HT #_ #_ #HT -elim (ffdeq_inv_gen_sn … HU) -HU #_ #_ #HU -/3 width=5 by tdeq_repl/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/ffdeq_fqup.ma b/matita/matita/contribs/lambdadelta/basic_2/static/ffdeq_fqup.ma deleted file mode 100644 index 15baed91f..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/ffdeq_fqup.ma +++ /dev/null @@ -1,29 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfdeq_fqup.ma". -include "basic_2/static/ffdeq.ma". - -(* DEGREE-BASED EQUIVALENCE FOR CLOSURES ON REFERRED ENTRIES ****************) - -(* Properties with degree-based equivalence for terms ***********************) - -lemma tdeq_ffdeq: ∀h,o,T1,T2. T1 ≛[h, o] T2 → - ∀G,L. ⦃G, L, T1⦄ ≛[h, o] ⦃G, L, T2⦄. -/2 width=1 by ffdeq_intro_sn/ qed. - -(* Advanced properties ******************************************************) - -lemma ffdeq_refl: ∀h,o. tri_reflexive … (ffdeq h o). -/2 width=1 by ffdeq_intro_sn/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/ffdeq_fqus.ma b/matita/matita/contribs/lambdadelta/basic_2/static/ffdeq_fqus.ma deleted file mode 100644 index 46a9640d9..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/ffdeq_fqus.ma +++ /dev/null @@ -1,31 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfdeq_fqus.ma". -include "basic_2/static/ffdeq.ma". - -(* DEGREE-BASED EQUIVALENCE FOR CLOSURES ON REFERRED ENTRIES ****************) - -(* Properties with star-iterated structural successor for closures **********) - -lemma ffdeq_fqus_trans: ∀h,o,b,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G, L, T⦄ → - ∀G2,L2,T2. ⦃G, L, T⦄ ⊐*[b] ⦃G2, L2, T2⦄ → - ∃∃G,L0,T0. ⦃G1, L1, T1⦄ ⊐*[b] ⦃G, L0, T0⦄ & ⦃G, L0, T0⦄ ≛[h, o] ⦃G2, L2, T2⦄. -#h #o #b #G1 #G #L1 #L #T1 #T #H1 #G2 #L2 #T2 #H2 -elim(ffdeq_inv_gen_dx … H1) -H1 #HG #HL1 #HT1 destruct -elim (lfdeq_fqus_trans … H2 … HL1) -L #L #T0 #H2 #HT02 #HL2 -elim (tdeq_fqus_trans … H2 … HT1) -T #L0 #T #H2 #HT0 #HL0 -lapply (tdeq_lfdeq_conf … HT02 … HL0) -HL0 #HL0 -/4 width=7 by ffdeq_intro_dx, lfdeq_trans, tdeq_trans, ex2_3_intro/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/ffdeq_lfeq.ma b/matita/matita/contribs/lambdadelta/basic_2/static/ffdeq_lfeq.ma deleted file mode 100644 index a0bf7c46d..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/ffdeq_lfeq.ma +++ /dev/null @@ -1,27 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfdeq_lfeq.ma". -include "basic_2/static/ffdeq.ma". - -(* DEGREE-BASED EQUIVALENCE FOR CLOSURES ON REFERRED ENTRIES ****************) - -(* Properties with syntactic equivalence on referred entries ****************) - -lemma lfeq_lfdeq_trans: ∀h,o,L1,L,T1. L1 ≡[T1] L → - ∀G1,G2,L2,T2. ⦃G1, L, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄. -#h #o #L1 #L #T1 #HL1 #G1 #G2 #L2 #T2 #H -elim (ffdeq_inv_gen_sn … H) -H #H #HL2 #T12 destruct -/3 width=3 by ffdeq_intro_sn, lfeq_lfdeq_trans/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma index 75e693912..fe24eadf5 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma @@ -191,7 +191,7 @@ lemma frees_inv_lifts: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≘ f2 → /3 width=7 by frees_eq_repl_back, coafter_inj/ qed-. -(* Note: this is used by lfxs_conf and might be modified *) +(* Note: this is used by rex_conf and might be modified *) lemma frees_inv_drops_next: ∀f1,L1,T1. L1 ⊢ 𝐅*⦃T1⦄ ≘ f1 → ∀I2,L2,V2,n. ⬇*[n] L1 ≘ L2.ⓑ{I2}V2 → ∀g1. ↑g1 = ⫱*[n] f1 → diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq.ma deleted file mode 100644 index 4103c54c5..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq.ma +++ /dev/null @@ -1,194 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/notation/relations/stareqsn_5.ma". -include "basic_2/syntax/tdeq_ext.ma". -include "basic_2/static/lfxs.ma". - -(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******) - -definition lfdeq (h) (o): relation3 term lenv lenv ≝ - lfxs (cdeq h o). - -interpretation - "degree-based equivalence on referred entries (local environment)" - 'StarEqSn h o T L1 L2 = (lfdeq h o T L1 L2). - -interpretation - "degree-based ranged equivalence (local environment)" - 'StarEqSn h o f L1 L2 = (lexs (cdeq_ext h o) cfull f L1 L2). - -(* Basic properties ***********************************************************) - -lemma frees_tdeq_conf_lfdeq (h) (o): ∀f,L1,T1. L1 ⊢ 𝐅*⦃T1⦄ ≘ f → ∀T2. T1 ≛[h, o] T2 → - ∀L2. L1 ≛[h, o, f] L2 → L2 ⊢ 𝐅*⦃T2⦄ ≘ f. -#h #o #f #L1 #T1 #H elim H -f -L1 -T1 -[ #f #L1 #s1 #Hf #X #H1 #L2 #_ - elim (tdeq_inv_sort1 … H1) -H1 #s2 #d #_ #_ #H destruct - /2 width=3 by frees_sort/ -| #f #i #Hf #X #H1 - >(tdeq_inv_lref1 … H1) -X #Y #H2 - >(lexs_inv_atom1 … H2) -Y - /2 width=1 by frees_atom/ -| #f #I #L1 #V1 #_ #IH #X #H1 - >(tdeq_inv_lref1 … H1) -X #Y #H2 - elim (lexs_inv_next1 … H2) -H2 #Z #L2 #HL12 #HZ #H destruct - elim (ext2_inv_pair_sn … HZ) -HZ #V2 #HV12 #H destruct - /3 width=1 by frees_pair/ -| #f #I #L1 #Hf #X #H1 - >(tdeq_inv_lref1 … H1) -X #Y #H2 - elim (lexs_inv_next1 … H2) -H2 #Z #L2 #_ #HZ #H destruct - >(ext2_inv_unit_sn … HZ) -Z /2 width=1 by frees_unit/ -| #f #I #L1 #i #_ #IH #X #H1 - >(tdeq_inv_lref1 … H1) -X #Y #H2 - elim (lexs_inv_push1 … H2) -H2 #J #L2 #HL12 #_ #H destruct - /3 width=1 by frees_lref/ -| #f #L1 #l #Hf #X #H1 #L2 #_ - >(tdeq_inv_gref1 … H1) -X /2 width=1 by frees_gref/ -| #f1V #f1T #f1 #p #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #X #H1 - elim (tdeq_inv_pair1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H1 #L2 #HL12 destruct - /6 width=5 by frees_bind, lexs_inv_tl, ext2_pair, sle_lexs_trans, sor_inv_sle_dx, sor_inv_sle_sn/ -| #f1V #f1T #f1 #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #X #H1 - elim (tdeq_inv_pair1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H1 #L2 #HL12 destruct - /5 width=5 by frees_flat, sle_lexs_trans, sor_inv_sle_dx, sor_inv_sle_sn/ -] -qed-. - -lemma frees_tdeq_conf (h) (o): ∀f,L,T1. L ⊢ 𝐅*⦃T1⦄ ≘ f → - ∀T2. T1 ≛[h, o] T2 → L ⊢ 𝐅*⦃T2⦄ ≘ f. -/4 width=7 by frees_tdeq_conf_lfdeq, lexs_refl, ext2_refl/ qed-. - -lemma frees_lfdeq_conf (h) (o): ∀f,L1,T. L1 ⊢ 𝐅*⦃T⦄ ≘ f → - ∀L2. L1 ≛[h, o, f] L2 → L2 ⊢ 𝐅*⦃T⦄ ≘ f. -/2 width=7 by frees_tdeq_conf_lfdeq, tdeq_refl/ qed-. - -lemma tdeq_lfxs_conf (R) (h) (o): s_r_confluent1 … (cdeq h o) (lfxs R). -#R #h #o #L1 #T1 #T2 #HT12 #L2 * -/3 width=5 by frees_tdeq_conf, ex2_intro/ -qed-. - -lemma tdeq_lfxs_div (R) (h) (o): ∀T1,T2. T1 ≛[h, o] T2 → - ∀L1,L2. L1 ⪤*[R, T2] L2 → L1 ⪤*[R, T1] L2. -/3 width=5 by tdeq_lfxs_conf, tdeq_sym/ qed-. - -lemma tdeq_lfdeq_conf (h) (o): s_r_confluent1 … (cdeq h o) (lfdeq h o). -/2 width=5 by tdeq_lfxs_conf/ qed-. - -lemma tdeq_lfdeq_div (h) (o): ∀T1,T2. T1 ≛[h, o] T2 → - ∀L1,L2. L1 ≛[h, o, T2] L2 → L1 ≛[h, o, T1] L2. -/2 width=5 by tdeq_lfxs_div/ qed-. - -lemma lfdeq_atom (h) (o): ∀I. ⋆ ≛[h, o, ⓪{I}] ⋆. -/2 width=1 by lfxs_atom/ qed. - -lemma lfdeq_sort (h) (o): ∀I1,I2,L1,L2,s. - L1 ≛[h, o, ⋆s] L2 → L1.ⓘ{I1} ≛[h, o, ⋆s] L2.ⓘ{I2}. -/2 width=1 by lfxs_sort/ qed. - -lemma lfdeq_pair (h) (o): ∀I,L1,L2,V1,V2. L1 ≛[h, o, V1] L2 → V1 ≛[h, o] V2 → - L1.ⓑ{I}V1 ≛[h, o, #0] L2.ⓑ{I}V2. -/2 width=1 by lfxs_pair/ qed. -(* -lemma lfdeq_unit (h) (o): ∀f,I,L1,L2. 𝐈⦃f⦄ → L1 ⪤*[cdeq_ext h o, cfull, f] L2 → - L1.ⓤ{I} ≛[h, o, #0] L2.ⓤ{I}. -/2 width=3 by lfxs_unit/ qed. -*) -lemma lfdeq_lref (h) (o): ∀I1,I2,L1,L2,i. - L1 ≛[h, o, #i] L2 → L1.ⓘ{I1} ≛[h, o, #↑i] L2.ⓘ{I2}. -/2 width=1 by lfxs_lref/ qed. - -lemma lfdeq_gref (h) (o): ∀I1,I2,L1,L2,l. - L1 ≛[h, o, §l] L2 → L1.ⓘ{I1} ≛[h, o, §l] L2.ⓘ{I2}. -/2 width=1 by lfxs_gref/ qed. - -lemma lfdeq_bind_repl_dx (h) (o): ∀I,I1,L1,L2.∀T:term. - L1.ⓘ{I} ≛[h, o, T] L2.ⓘ{I1} → - ∀I2. I ≛[h, o] I2 → - L1.ⓘ{I} ≛[h, o, T] L2.ⓘ{I2}. -/2 width=2 by lfxs_bind_repl_dx/ qed-. - -(* Basic inversion lemmas ***************************************************) - -lemma lfdeq_inv_atom_sn (h) (o): ∀Y2. ∀T:term. ⋆ ≛[h, o, T] Y2 → Y2 = ⋆. -/2 width=3 by lfxs_inv_atom_sn/ qed-. - -lemma lfdeq_inv_atom_dx (h) (o): ∀Y1. ∀T:term. Y1 ≛[h, o, T] ⋆ → Y1 = ⋆. -/2 width=3 by lfxs_inv_atom_dx/ qed-. -(* -lemma lfdeq_inv_zero (h) (o): ∀Y1,Y2. Y1 ≛[h, o, #0] Y2 → - ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆ - | ∃∃I,L1,L2,V1,V2. L1 ≛[h, o, V1] L2 & V1 ≛[h, o] V2 & - Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2 - | ∃∃f,I,L1,L2. 𝐈⦃f⦄ & L1 ⪤*[cdeq_ext h o, cfull, f] L2 & - Y1 = L1.ⓤ{I} & Y2 = L2.ⓤ{I}. -#h #o #Y1 #Y2 #H elim (lfxs_inv_zero … H) -H * -/3 width=9 by or3_intro0, or3_intro1, or3_intro2, ex4_5_intro, ex4_4_intro, conj/ -qed-. -*) -lemma lfdeq_inv_lref (h) (o): ∀Y1,Y2,i. Y1 ≛[h, o, #↑i] Y2 → - ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆ - | ∃∃I1,I2,L1,L2. L1 ≛[h, o, #i] L2 & - Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. -/2 width=1 by lfxs_inv_lref/ qed-. - -(* Basic_2A1: uses: lleq_inv_bind lleq_inv_bind_O *) -lemma lfdeq_inv_bind (h) (o): ∀p,I,L1,L2,V,T. L1 ≛[h, o, ⓑ{p,I}V.T] L2 → - ∧∧ L1 ≛[h, o, V] L2 & L1.ⓑ{I}V ≛[h, o, T] L2.ⓑ{I}V. -/2 width=2 by lfxs_inv_bind/ qed-. - -(* Basic_2A1: uses: lleq_inv_flat *) -lemma lfdeq_inv_flat (h) (o): ∀I,L1,L2,V,T. L1 ≛[h, o, ⓕ{I}V.T] L2 → - ∧∧ L1 ≛[h, o, V] L2 & L1 ≛[h, o, T] L2. -/2 width=2 by lfxs_inv_flat/ qed-. - -(* Advanced inversion lemmas ************************************************) - -lemma lfdeq_inv_zero_pair_sn (h) (o): ∀I,Y2,L1,V1. L1.ⓑ{I}V1 ≛[h, o, #0] Y2 → - ∃∃L2,V2. L1 ≛[h, o, V1] L2 & V1 ≛[h, o] V2 & Y2 = L2.ⓑ{I}V2. -/2 width=1 by lfxs_inv_zero_pair_sn/ qed-. - -lemma lfdeq_inv_zero_pair_dx (h) (o): ∀I,Y1,L2,V2. Y1 ≛[h, o, #0] L2.ⓑ{I}V2 → - ∃∃L1,V1. L1 ≛[h, o, V1] L2 & V1 ≛[h, o] V2 & Y1 = L1.ⓑ{I}V1. -/2 width=1 by lfxs_inv_zero_pair_dx/ qed-. - -lemma lfdeq_inv_lref_bind_sn (h) (o): ∀I1,Y2,L1,i. L1.ⓘ{I1} ≛[h, o, #↑i] Y2 → - ∃∃I2,L2. L1 ≛[h, o, #i] L2 & Y2 = L2.ⓘ{I2}. -/2 width=2 by lfxs_inv_lref_bind_sn/ qed-. - -lemma lfdeq_inv_lref_bind_dx (h) (o): ∀I2,Y1,L2,i. Y1 ≛[h, o, #↑i] L2.ⓘ{I2} → - ∃∃I1,L1. L1 ≛[h, o, #i] L2 & Y1 = L1.ⓘ{I1}. -/2 width=2 by lfxs_inv_lref_bind_dx/ qed-. - -(* Basic forward lemmas *****************************************************) - -lemma lfdeq_fwd_zero_pair (h) (o): ∀I,K1,K2,V1,V2. - K1.ⓑ{I}V1 ≛[h, o, #0] K2.ⓑ{I}V2 → K1 ≛[h, o, V1] K2. -/2 width=3 by lfxs_fwd_zero_pair/ qed-. - -(* Basic_2A1: uses: lleq_fwd_bind_sn lleq_fwd_flat_sn *) -lemma lfdeq_fwd_pair_sn (h) (o): ∀I,L1,L2,V,T. L1 ≛[h, o, ②{I}V.T] L2 → L1 ≛[h, o, V] L2. -/2 width=3 by lfxs_fwd_pair_sn/ qed-. - -(* Basic_2A1: uses: lleq_fwd_bind_dx lleq_fwd_bind_O_dx *) -lemma lfdeq_fwd_bind_dx (h) (o): ∀p,I,L1,L2,V,T. - L1 ≛[h, o, ⓑ{p,I}V.T] L2 → L1.ⓑ{I}V ≛[h, o, T] L2.ⓑ{I}V. -/2 width=2 by lfxs_fwd_bind_dx/ qed-. - -(* Basic_2A1: uses: lleq_fwd_flat_dx *) -lemma lfdeq_fwd_flat_dx (h) (o): ∀I,L1,L2,V,T. L1 ≛[h, o, ⓕ{I}V.T] L2 → L1 ≛[h, o, T] L2. -/2 width=3 by lfxs_fwd_flat_dx/ qed-. - -lemma lfdeq_fwd_dx (h) (o): ∀I2,L1,K2. ∀T:term. L1 ≛[h, o, T] K2.ⓘ{I2} → - ∃∃I1,K1. L1 = K1.ⓘ{I1}. -/2 width=5 by lfxs_fwd_dx/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_drops.ma deleted file mode 100644 index 8e0e253e0..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_drops.ma +++ /dev/null @@ -1,47 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/relocation/lifts_tdeq.ma". -include "basic_2/static/lfxs_drops.ma". -include "basic_2/static/lfdeq.ma". - -(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******) - -(* Properties with generic slicing for local environments *******************) - -lemma lfdeq_lifts_sn: ∀h,o. f_dedropable_sn (cdeq h o). -/3 width=5 by lfxs_liftable_dedropable_sn, tdeq_lifts_sn/ qed-. - -(* Inversion lemmas with generic slicing for local environments *************) - -lemma lfdeq_inv_lifts_sn: ∀h,o. f_dropable_sn (cdeq h o). -/2 width=5 by lfxs_dropable_sn/ qed-. - -(* Note: missing in basic_2A1 *) -lemma lfdeq_inv_lifts_dx: ∀h,o. f_dropable_dx (cdeq h o). -/2 width=5 by lfxs_dropable_dx/ qed-. - -(* Basic_2A1: uses: lleq_inv_lift_le lleq_inv_lift_be lleq_inv_lift_ge *) -lemma lfdeq_inv_lifts_bi: ∀h,o,L1,L2,U. L1 ≛[h, o, U] L2 → ∀b,f. 𝐔⦃f⦄ → - ∀K1,K2. ⬇*[b, f] L1 ≘ K1 → ⬇*[b, f] L2 ≘ K2 → - ∀T. ⬆*[f] T ≘ U → K1 ≛[h, o, T] K2. -/2 width=10 by lfxs_inv_lifts_bi/ qed-. - -lemma lfdeq_inv_lref_pair_sn: ∀h,o,L1,L2,i. L1 ≛[h, o, #i] L2 → ∀I,K1,V1. ⬇*[i] L1 ≘ K1.ⓑ{I}V1 → - ∃∃K2,V2. ⬇*[i] L2 ≘ K2.ⓑ{I}V2 & K1 ≛[h, o, V1] K2 & V1 ≛[h, o] V2. -/2 width=3 by lfxs_inv_lref_pair_sn/ qed-. - -lemma lfdeq_inv_lref_pair_dx: ∀h,o,L1,L2,i. L1 ≛[h, o, #i] L2 → ∀I,K2,V2. ⬇*[i] L2 ≘ K2.ⓑ{I}V2 → - ∃∃K1,V1. ⬇*[i] L1 ≘ K1.ⓑ{I}V1 & K1 ≛[h, o, V1] K2 & V1 ≛[h, o] V2. -/2 width=3 by lfxs_inv_lref_pair_dx/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_fqup.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_fqup.ma deleted file mode 100644 index c39dd5bbc..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_fqup.ma +++ /dev/null @@ -1,39 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfxs_fqup.ma". -include "basic_2/static/lfdeq.ma". - -(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******) - -(* Advanced properties ******************************************************) - -lemma lfdeq_refl: ∀h,o,T. reflexive … (lfdeq h o T). -/2 width=1 by lfxs_refl/ qed. - -lemma lfdeq_pair_refl: ∀h,o,V1,V2. V1 ≛[h, o] V2 → - ∀I,L. ∀T:term. L.ⓑ{I}V1 ≛[h, o, T] L.ⓑ{I}V2. -/2 width=1 by lfxs_pair_refl/ qed. - -(* Advanced inversion lemmas ************************************************) - -lemma lfdeq_inv_bind_void: ∀h,o,p,I,L1,L2,V,T. L1 ≛[h, o, ⓑ{p,I}V.T] L2 → - L1 ≛[h, o, V] L2 ∧ L1.ⓧ ≛[h, o, T] L2.ⓧ. -/2 width=3 by lfxs_inv_bind_void/ qed-. - -(* Advanced forward lemmas **************************************************) - -lemma lfdeq_fwd_bind_dx_void: ∀h,o,p,I,L1,L2,V,T. - L1 ≛[h, o, ⓑ{p,I}V.T] L2 → L1.ⓧ ≛[h, o, T] L2.ⓧ. -/2 width=4 by lfxs_fwd_bind_dx_void/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_fqus.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_fqus.ma deleted file mode 100644 index 84685ca10..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_fqus.ma +++ /dev/null @@ -1,156 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/s_computation/fqus_fqup.ma". -include "basic_2/static/lfdeq_drops.ma". -include "basic_2/static/lfdeq_fqup.ma". -include "basic_2/static/lfdeq_lfdeq.ma". - -(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******) - -(* Properties with extended structural successor for closures ***************) - -lemma fqu_tdeq_conf: ∀h,o,b,G1,G2,L1,L2,U1,T1. ⦃G1, L1, U1⦄ ⊐[b] ⦃G2, L2, T1⦄ → - ∀U2. U1 ≛[h, o] U2 → - ∃∃L,T2. ⦃G1, L1, U2⦄ ⊐[b] ⦃G2, L, T2⦄ & L2 ≛[h, o, T1] L & T1 ≛[h, o] T2. -#h #o #b #G1 #G2 #L1 #L2 #U1 #T1 #H elim H -G1 -G2 -L1 -L2 -U1 -T1 -[ #I #G #L #W #X #H >(tdeq_inv_lref1 … H) -X - /2 width=5 by fqu_lref_O, ex3_2_intro/ -| #I #G #L #W1 #U1 #X #H - elim (tdeq_inv_pair1 … H) -H #W2 #U2 #HW12 #_ #H destruct - /2 width=5 by fqu_pair_sn, ex3_2_intro/ -| #p #I #G #L #W1 #U1 #X #H - elim (tdeq_inv_pair1 … H) -H #W2 #U2 #HW12 #HU12 #H destruct - /3 width=5 by lfdeq_pair_refl, fqu_bind_dx, ex3_2_intro/ -| #p #I #G #L #W1 #U1 #Hb #X #H - elim (tdeq_inv_pair1 … H) -H #W2 #U2 #HW12 #HU12 #H destruct - /3 width=5 by fqu_clear, ex3_2_intro/ -| #I #G #L #W1 #U1 #X #H - elim (tdeq_inv_pair1 … H) -H #W2 #U2 #_ #HU12 #H destruct - /2 width=5 by fqu_flat_dx, ex3_2_intro/ -| #I #G #L #T1 #U1 #HTU1 #U2 #HU12 - elim (tdeq_inv_lifts_sn … HU12 … HTU1) -U1 - /3 width=5 by fqu_drop, ex3_2_intro/ -] -qed-. - -lemma tdeq_fqu_trans: ∀h,o,b,G1,G2,L1,L2,U1,T1. ⦃G1, L1, U1⦄ ⊐[b] ⦃G2, L2, T1⦄ → - ∀U2. U2 ≛[h, o] U1 → - ∃∃L,T2. ⦃G1, L1, U2⦄ ⊐[b] ⦃G2, L, T2⦄ & T2 ≛[h, o] T1 & L ≛[h, o, T1] L2. -#h #o #b #G1 #G2 #L1 #L2 #U1 #T1 #H12 #U2 #HU21 -elim (fqu_tdeq_conf … o … H12 U2) -H12 -/3 width=5 by lfdeq_sym, tdeq_sym, ex3_2_intro/ -qed-. - -(* Basic_2A1: uses: lleq_fqu_trans *) -lemma lfdeq_fqu_trans: ∀h,o,b,G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊐[b] ⦃G2, K2, U⦄ → - ∀L1. L1 ≛[h, o, T] L2 → - ∃∃K1,U0. ⦃G1, L1, T⦄ ⊐[b] ⦃G2, K1, U0⦄ & U0 ≛[h, o] U & K1 ≛[h, o, U] K2. -#h #o #b #G1 #G2 #L2 #K2 #T #U #H elim H -G1 -G2 -L2 -K2 -T -U -[ #I #G #L2 #V2 #L1 #H elim (lfdeq_inv_zero_pair_dx … H) -H - #K1 #V1 #HV1 #HV12 #H destruct - /3 width=7 by tdeq_lfdeq_conf, fqu_lref_O, ex3_2_intro/ -| * [ #p ] #I #G #L2 #V #T #L1 #H - [ elim (lfdeq_inv_bind … H) - | elim (lfdeq_inv_flat … H) - ] -H - /2 width=5 by fqu_pair_sn, ex3_2_intro/ -| #p #I #G #L2 #V #T #L1 #H elim (lfdeq_inv_bind … H) -H - /2 width=5 by fqu_bind_dx, ex3_2_intro/ -| #p #I #G #L2 #V #T #Hb #L1 #H elim (lfdeq_inv_bind_void … H) -H - /3 width=5 by fqu_clear, ex3_2_intro/ -| #I #G #L2 #V #T #L1 #H elim (lfdeq_inv_flat … H) -H - /2 width=5 by fqu_flat_dx, ex3_2_intro/ -| #I #G #L2 #T #U #HTU #Y #HU - elim (lfdeq_fwd_dx … HU) #L1 #V1 #H destruct - /5 width=14 by lfdeq_inv_lifts_bi, fqu_drop, drops_refl, drops_drop, ex3_2_intro/ -] -qed-. - -(* Properties with optional structural successor for closures ***************) - -lemma tdeq_fquq_trans: ∀h,o,b,G1,G2,L1,L2,U1,T1. ⦃G1, L1, U1⦄ ⊐⸮[b] ⦃G2, L2, T1⦄ → - ∀U2. U2 ≛[h, o] U1 → - ∃∃L,T2. ⦃G1, L1, U2⦄ ⊐⸮[b] ⦃G2, L, T2⦄ & T2 ≛[h, o] T1 & L ≛[h, o, T1] L2. -#h #o #b #G1 #G2 #L1 #L2 #U1 #T1 #H elim H -H -[ #H #U2 #HU21 elim (tdeq_fqu_trans … H … HU21) -U1 - /3 width=5 by fqu_fquq, ex3_2_intro/ -| * #HG #HL #HT destruct /2 width=5 by ex3_2_intro/ -] -qed-. - -(* Basic_2A1: was just: lleq_fquq_trans *) -lemma lfdeq_fquq_trans: ∀h,o,b,G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊐⸮[b] ⦃G2, K2, U⦄ → - ∀L1. L1 ≛[h, o, T] L2 → - ∃∃K1,U0. ⦃G1, L1, T⦄ ⊐⸮[b] ⦃G2, K1, U0⦄ & U0 ≛[h, o] U & K1 ≛[h, o, U] K2. -#h #o #b #G1 #G2 #L2 #K2 #T #U #H elim H -H -[ #H #L1 #HL12 elim (lfdeq_fqu_trans … H … HL12) -L2 /3 width=5 by fqu_fquq, ex3_2_intro/ -| * #HG #HL #HT destruct /2 width=5 by ex3_2_intro/ -] -qed-. - -(* Properties with plus-iterated structural successor for closures **********) - -(* Basic_2A1: was just: lleq_fqup_trans *) -lemma lfdeq_fqup_trans: ∀h,o,b,G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊐+[b] ⦃G2, K2, U⦄ → - ∀L1. L1 ≛[h, o, T] L2 → - ∃∃K1,U0. ⦃G1, L1, T⦄ ⊐+[b] ⦃G2, K1, U0⦄ & U0 ≛[h, o] U & K1 ≛[h, o, U] K2. -#h #o #b #G1 #G2 #L2 #K2 #T #U #H @(fqup_ind … H) -G2 -K2 -U -[ #G2 #K2 #U #HTU #L1 #HL12 elim (lfdeq_fqu_trans … HTU … HL12) -L2 - /3 width=5 by fqu_fqup, ex3_2_intro/ -| #G #G2 #K #K2 #U #U2 #_ #HU2 #IHTU #L1 #HL12 - elim (IHTU … HL12) -L2 #K0 #U0 #HTU #HU0 #HK0 - elim (lfdeq_fqu_trans … HU2 … HK0) -K #K1 #U1 #HU1 #HU12 #HK12 - elim (tdeq_fqu_trans … HU1 … HU0) -U #K3 #U3 #HU03 #HU31 #HK31 - @(ex3_2_intro … K3 U3) (**) (* full auto too slow *) - /3 width=5 by lfdeq_trans, tdeq_lfdeq_conf, fqup_strap1, tdeq_trans/ -] -qed-. - -lemma tdeq_fqup_trans: ∀h,o,b,G1,G2,L1,L2,U1,T1. ⦃G1, L1, U1⦄ ⊐+[b] ⦃G2, L2, T1⦄ → - ∀U2. U2 ≛[h, o] U1 → - ∃∃L,T2. ⦃G1, L1, U2⦄ ⊐+[b] ⦃G2, L, T2⦄ & T2 ≛[h, o] T1 & L ≛[h, o, T1] L2. -#h #o #b #G1 #G2 #L1 #L2 #U1 #T1 #H @(fqup_ind_dx … H) -G1 -L1 -U1 -[ #G1 #L1 #U1 #H #U2 #HU21 elim (tdeq_fqu_trans … H … HU21) -U1 - /3 width=5 by fqu_fqup, ex3_2_intro/ -| #G1 #G #L1 #L #U1 #U #H #_ #IH #U2 #HU21 - elim (tdeq_fqu_trans … H … HU21) -U1 #L0 #T #H1 #HTU #HL0 - lapply (tdeq_lfdeq_div … HTU … HL0) -HL0 #HL0 - elim (IH … HTU) -U #K2 #U1 #H2 #HUT1 #HKL2 - elim (lfdeq_fqup_trans … H2 … HL0) -L #K #U #H2 #HU1 #HK2 - lapply (tdeq_lfdeq_conf … HUT1 … HK2) -HK2 #HK2 - /3 width=7 by lfdeq_trans, fqup_strap2, tdeq_trans, ex3_2_intro/ -] -qed-. - -(* Properties with star-iterated structural successor for closures **********) - -lemma tdeq_fqus_trans: ∀h,o,b,G1,G2,L1,L2,U1,T1. ⦃G1, L1, U1⦄ ⊐*[b] ⦃G2, L2, T1⦄ → - ∀U2. U2 ≛[h, o] U1 → - ∃∃L,T2. ⦃G1, L1, U2⦄ ⊐*[b] ⦃G2, L, T2⦄ & T2 ≛[h, o] T1 & L ≛[h, o, T1] L2. -#h #o #b #G1 #G2 #L1 #L2 #U1 #T1 #H #U2 #HU21 elim(fqus_inv_fqup … H) -H -[ #H elim (tdeq_fqup_trans … H … HU21) -U1 /3 width=5 by fqup_fqus, ex3_2_intro/ -| * #HG #HL #HT destruct /2 width=5 by ex3_2_intro/ -] -qed-. - -(* Basic_2A1: was just: lleq_fqus_trans *) -lemma lfdeq_fqus_trans: ∀h,o,b,G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊐*[b] ⦃G2, K2, U⦄ → - ∀L1. L1 ≛[h, o, T] L2 → - ∃∃K1,U0. ⦃G1, L1, T⦄ ⊐*[b] ⦃G2, K1, U0⦄ & U0 ≛[h, o] U & K1 ≛[h, o, U] K2. -#h #o #b #G1 #G2 #L2 #K2 #T #U #H #L1 #HL12 elim(fqus_inv_fqup … H) -H -[ #H elim (lfdeq_fqup_trans … H … HL12) -L2 /3 width=5 by fqup_fqus, ex3_2_intro/ -| * #HG #HL #HT destruct /2 width=5 by ex3_2_intro/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_length.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_length.ma deleted file mode 100644 index 40f2812cc..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_length.ma +++ /dev/null @@ -1,55 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/relocation/lifts_tdeq.ma". -include "basic_2/static/lfxs_length.ma". -include "basic_2/static/lfxs_fsle.ma". -include "basic_2/static/lfdeq.ma". - -(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******) - -(* Advanved properties with free variables inclusion ************************) - -lemma lfdeq_fsge_comp (h) (o): lfxs_fsge_compatible (cdeq h o). -#h #o #L1 #L2 #T * #f1 #Hf1 #HL12 -lapply (frees_lfdeq_conf h o … Hf1 … HL12) -lapply (lexs_fwd_length … HL12) -/3 width=8 by lveq_length_eq, ex4_4_intro/ (**) (* full auto fails *) -qed-. - -(* Properties with length for local environments ****************************) - -(* Basic_2A1: uses: lleq_sort *) -lemma lfdeq_sort_length (h) (o): ∀L1,L2. |L1| = |L2| → ∀s. L1 ≛[h, o, ⋆s] L2. -/2 width=1 by lfxs_sort_length/ qed. - -(* Basic_2A1: uses: lleq_gref *) -lemma lfdeq_gref_length (h) (o): ∀L1,L2. |L1| = |L2| → ∀l. L1 ≛[h, o, §l] L2. -/2 width=1 by lfxs_gref_length/ qed. - -lemma lfdeq_unit_length (h) (o): ∀L1,L2. |L1| = |L2| → - ∀I. L1.ⓤ{I} ≛[h, o, #0] L2.ⓤ{I}. -/2 width=1 by lfxs_unit_length/ qed. - -(* Basic_2A1: uses: lleq_lift_le lleq_lift_ge *) -lemma lfdeq_lifts_bi (h) (o): ∀L1,L2. |L1| = |L2| → ∀K1,K2,T. K1 ≛[h, o, T] K2 → - ∀b,f. ⬇*[b, f] L1 ≘ K1 → ⬇*[b, f] L2 ≘ K2 → - ∀U. ⬆*[f] T ≘ U → L1 ≛[h, o, U] L2. -/3 width=9 by lfxs_lifts_bi, tdeq_lifts_sn/ qed-. - -(* Forward lemmas with length for local environments ************************) - -(* Basic_2A1: lleq_fwd_length *) -lemma lfdeq_fwd_length (h) (o): ∀L1,L2. ∀T:term. L1 ≛[h, o, T] L2 → |L1| = |L2|. -/2 width=3 by lfxs_fwd_length/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_lfdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_lfdeq.ma deleted file mode 100644 index 3dab077cf..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_lfdeq.ma +++ /dev/null @@ -1,100 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/syntax/ext2_ext2.ma". -include "basic_2/syntax/tdeq_tdeq.ma". -include "basic_2/static/lfdeq_length.ma". - -(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******) - -(* Advanced properties ******************************************************) - -(* Basic_2A1: uses: lleq_sym *) -lemma lfdeq_sym: ∀h,o,T. symmetric … (lfdeq h o T). -/3 width=3 by lfdeq_fsge_comp, lfxs_sym, tdeq_sym/ qed-. - -(* Basic_2A1: uses: lleq_dec *) -lemma lfdeq_dec: ∀h,o,L1,L2. ∀T:term. Decidable (L1 ≛[h, o, T] L2). -/3 width=1 by lfxs_dec, tdeq_dec/ qed-. - -(* Main properties **********************************************************) - -(* Basic_2A1: uses: lleq_bind lleq_bind_O *) -theorem lfdeq_bind: ∀h,o,p,I,L1,L2,V1,V2,T. - L1 ≛[h, o, V1] L2 → L1.ⓑ{I}V1 ≛[h, o, T] L2.ⓑ{I}V2 → - L1 ≛[h, o, ⓑ{p,I}V1.T] L2. -/2 width=2 by lfxs_bind/ qed. - -(* Basic_2A1: uses: lleq_flat *) -theorem lfdeq_flat: ∀h,o,I,L1,L2,V,T. L1 ≛[h, o, V] L2 → L1 ≛[h, o, T] L2 → - L1 ≛[h, o, ⓕ{I}V.T] L2. -/2 width=1 by lfxs_flat/ qed. - -theorem lfdeq_bind_void: ∀h,o,p,I,L1,L2,V,T. - L1 ≛[h, o, V] L2 → L1.ⓧ ≛[h, o, T] L2.ⓧ → - L1 ≛[h, o, ⓑ{p,I}V.T] L2. -/2 width=1 by lfxs_bind_void/ qed. - -(* Basic_2A1: uses: lleq_trans *) -theorem lfdeq_trans: ∀h,o,T. Transitive … (lfdeq h o T). -#h #o #T #L1 #L * #f1 #Hf1 #HL1 #L2 * #f2 #Hf2 #HL2 -lapply (frees_tdeq_conf_lfdeq … Hf1 T … HL1) // #H0 -lapply (frees_mono … Hf2 … H0) -Hf2 -H0 -/5 width=7 by lexs_trans, lexs_eq_repl_back, tdeq_trans, ext2_trans, ex2_intro/ -qed-. - -(* Basic_2A1: uses: lleq_canc_sn *) -theorem lfdeq_canc_sn: ∀h,o,T. left_cancellable … (lfdeq h o T). -/3 width=3 by lfdeq_trans, lfdeq_sym/ qed-. - -(* Basic_2A1: uses: lleq_canc_dx *) -theorem lfdeq_canc_dx: ∀h,o,T. right_cancellable … (lfdeq h o T). -/3 width=3 by lfdeq_trans, lfdeq_sym/ qed-. - -theorem lfdeq_repl: ∀h,o,L1,L2. ∀T:term. L1 ≛[h, o, T] L2 → - ∀K1. L1 ≛[h, o, T] K1 → ∀K2. L2 ≛[h, o, T] K2 → K1 ≛[h, o, T] K2. -/3 width=3 by lfdeq_canc_sn, lfdeq_trans/ qed-. - -(* Negated properties *******************************************************) - -(* Note: auto works with /4 width=8/ so lfdeq_canc_sn is preferred **********) -(* Basic_2A1: uses: lleq_nlleq_trans *) -lemma lfdeq_lfdneq_trans: ∀h,o.∀T:term.∀L1,L. L1 ≛[h, o, T] L → - ∀L2. (L ≛[h, o, T] L2 → ⊥) → (L1 ≛[h, o, T] L2 → ⊥). -/3 width=3 by lfdeq_canc_sn/ qed-. - -(* Basic_2A1: uses: nlleq_lleq_div *) -lemma lfdneq_lfdeq_div: ∀h,o.∀T:term.∀L2,L. L2 ≛[h, o, T] L → - ∀L1. (L1 ≛[h, o, T] L → ⊥) → (L1 ≛[h, o, T] L2 → ⊥). -/3 width=3 by lfdeq_trans/ qed-. - -theorem lfdneq_lfdeq_canc_dx: ∀h,o,L1,L. ∀T:term. (L1 ≛[h, o, T] L → ⊥) → - ∀L2. L2 ≛[h, o, T] L → L1 ≛[h, o, T] L2 → ⊥. -/3 width=3 by lfdeq_trans/ qed-. - -(* Negated inversion lemmas *************************************************) - -(* Basic_2A1: uses: nlleq_inv_bind nlleq_inv_bind_O *) -lemma lfdneq_inv_bind: ∀h,o,p,I,L1,L2,V,T. (L1 ≛[h, o, ⓑ{p,I}V.T] L2 → ⊥) → - (L1 ≛[h, o, V] L2 → ⊥) ∨ (L1.ⓑ{I}V ≛[h, o, T] L2.ⓑ{I}V → ⊥). -/3 width=2 by lfnxs_inv_bind, tdeq_dec/ qed-. - -(* Basic_2A1: uses: nlleq_inv_flat *) -lemma lfdneq_inv_flat: ∀h,o,I,L1,L2,V,T. (L1 ≛[h, o, ⓕ{I}V.T] L2 → ⊥) → - (L1 ≛[h, o, V] L2 → ⊥) ∨ (L1 ≛[h, o, T] L2 → ⊥). -/3 width=2 by lfnxs_inv_flat, tdeq_dec/ qed-. - -lemma lfdneq_inv_bind_void: ∀h,o,p,I,L1,L2,V,T. (L1 ≛[h, o, ⓑ{p,I}V.T] L2 → ⊥) → - (L1 ≛[h, o, V] L2 → ⊥) ∨ (L1.ⓧ ≛[h, o, T] L2.ⓧ → ⊥). -/3 width=3 by lfnxs_inv_bind_void, tdeq_dec/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_lfeq.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_lfeq.ma deleted file mode 100644 index 39d933b76..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_lfeq.ma +++ /dev/null @@ -1,27 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfeq_fsle.ma". -include "basic_2/static/lfdeq.ma". - -(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******) - -(* Properties with syntactic equivalence on referred entries ****************) - -lemma lfeq_lfdeq: ∀h,o,L1,L2. ∀T:term. L1 ≡[T] L2 → L1 ≛[h, o, T] L2. -/2 width=3 by lfxs_co/ qed. - -lemma lfeq_lfdeq_trans: ∀h,o,L1,L. ∀T:term. L1 ≡[T] L → - ∀L2. L ≛[h, o, T] L2 → L1 ≛[h, o, T] L2. -/2 width=3 by lfeq_lfxs_trans/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfeq.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfeq.ma deleted file mode 100644 index f7375afbc..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfeq.ma +++ /dev/null @@ -1,109 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/notation/relations/ideqsn_3.ma". -include "basic_2/static/lfxs.ma". - -(* SYNTACTIC EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES *********) - -(* Basic_2A1: was: lleq *) -definition lfeq: relation3 term lenv lenv ≝ - lfxs ceq. - -interpretation - "syntactic equivalence on referred entries (local environment)" - 'IdEqSn T L1 L2 = (lfeq T L1 L2). - -(* Note: "lfeq_transitive R" is equivalent to "lfxs_transitive ceq R R" *) -(* Basic_2A1: uses: lleq_transitive *) -definition lfeq_transitive: predicate (relation3 lenv term term) ≝ - λR. ∀L2,T1,T2. R L2 T1 T2 → ∀L1. L1 ≡[T1] L2 → R L1 T1 T2. - -(* Basic inversion lemmas ***************************************************) - -lemma lfeq_inv_bind: ∀p,I,L1,L2,V,T. L1 ≡[ⓑ{p,I}V.T] L2 → - ∧∧ L1 ≡[V] L2 & L1.ⓑ{I}V ≡[T] L2.ⓑ{I}V. -/2 width=2 by lfxs_inv_bind/ qed-. - -lemma lfeq_inv_flat: ∀I,L1,L2,V,T. L1 ≡[ⓕ{I}V.T] L2 → - ∧∧ L1 ≡[V] L2 & L1 ≡[T] L2. -/2 width=2 by lfxs_inv_flat/ qed-. - -(* Advanced inversion lemmas ************************************************) - -lemma lfeq_inv_zero_pair_sn: ∀I,L2,K1,V. K1.ⓑ{I}V ≡[#0] L2 → - ∃∃K2. K1 ≡[V] K2 & L2 = K2.ⓑ{I}V. -#I #L2 #K1 #V #H -elim (lfxs_inv_zero_pair_sn … H) -H #K2 #X #HK12 #HX #H destruct -/2 width=3 by ex2_intro/ -qed-. - -lemma lfeq_inv_zero_pair_dx: ∀I,L1,K2,V. L1 ≡[#0] K2.ⓑ{I}V → - ∃∃K1. K1 ≡[V] K2 & L1 = K1.ⓑ{I}V. -#I #L1 #K2 #V #H -elim (lfxs_inv_zero_pair_dx … H) -H #K1 #X #HK12 #HX #H destruct -/2 width=3 by ex2_intro/ -qed-. - -lemma lfeq_inv_lref_bind_sn: ∀I1,K1,L2,i. K1.ⓘ{I1} ≡[#↑i] L2 → - ∃∃I2,K2. K1 ≡[#i] K2 & L2 = K2.ⓘ{I2}. -/2 width=2 by lfxs_inv_lref_bind_sn/ qed-. - -lemma lfeq_inv_lref_bind_dx: ∀I2,K2,L1,i. L1 ≡[#↑i] K2.ⓘ{I2} → - ∃∃I1,K1. K1 ≡[#i] K2 & L1 = K1.ⓘ{I1}. -/2 width=2 by lfxs_inv_lref_bind_dx/ qed-. - -(* Basic forward lemmas *****************************************************) - -(* Basic_2A1: was: llpx_sn_lrefl *) -(* Note: this should have been lleq_fwd_llpx_sn *) -lemma lfeq_fwd_lfxs: ∀R. c_reflexive … R → - ∀L1,L2,T. L1 ≡[T] L2 → L1 ⪤*[R, T] L2. -#R #HR #L1 #L2 #T * #f #Hf #HL12 -/4 width=7 by lexs_co, cext2_co, ex2_intro/ -qed-. - -(* Basic_properties *********************************************************) - -lemma frees_lfeq_conf: ∀f,L1,T. L1 ⊢ 𝐅*⦃T⦄ ≘ f → - ∀L2. L1 ≡[T] L2 → L2 ⊢ 𝐅*⦃T⦄ ≘ f. -#f #L1 #T #H elim H -f -L1 -T -[ /2 width=3 by frees_sort/ -| #f #i #Hf #L2 #H2 - >(lfxs_inv_atom_sn … H2) -L2 - /2 width=1 by frees_atom/ -| #f #I #L1 #V1 #_ #IH #Y #H2 - elim (lfeq_inv_zero_pair_sn … H2) -H2 #L2 #HL12 #H destruct - /3 width=1 by frees_pair/ -| #f #I #L1 #Hf #Y #H2 - elim (lfxs_inv_zero_unit_sn … H2) -H2 #g #L2 #_ #_ #H destruct - /2 width=1 by frees_unit/ -| #f #I #L1 #i #_ #IH #Y #H2 - elim (lfeq_inv_lref_bind_sn … H2) -H2 #J #L2 #HL12 #H destruct - /3 width=1 by frees_lref/ -| /2 width=1 by frees_gref/ -| #f1V #f1T #f1 #p #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #L2 #H2 - elim (lfeq_inv_bind … H2) -H2 /3 width=5 by frees_bind/ -| #f1V #f1T #f1 #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #L2 #H2 - elim (lfeq_inv_flat … H2) -H2 /3 width=5 by frees_flat/ -] -qed-. - -(* Basic_2A1: removed theorems 10: - lleq_ind lleq_fwd_lref - lleq_fwd_drop_sn lleq_fwd_drop_dx - lleq_skip lleq_lref lleq_free - lleq_Y lleq_ge_up lleq_ge - -*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfeq_fqup.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfeq_fqup.ma deleted file mode 100644 index 92479a038..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfeq_fqup.ma +++ /dev/null @@ -1,24 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfxs_fqup.ma". -include "basic_2/static/lfeq.ma". - -(* SYNTACTIC EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES *********) - -(* Advanced properties ******************************************************) - -(* Basic_2A1: was: lleq_refl *) -lemma lfeq_refl: ∀T. reflexive … (lfeq T). -/2 width=1 by lfxs_refl/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfeq_fsle.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfeq_fsle.ma deleted file mode 100644 index 0f80c9772..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfeq_fsle.ma +++ /dev/null @@ -1,33 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lfxs_length.ma". -include "basic_2/static/lfxs_fsle.ma". -include "basic_2/static/lfeq.ma". - -(* SYNTACTIC EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES *********) - -(* Properties with free variables inclusion for restricted closures *********) - -lemma lfeq_fsle_comp: lfxs_fsle_compatible ceq. -#L1 #L2 #T #HL12 -elim (frees_total L1 T) -/4 width=8 by frees_lfeq_conf, lfxs_fwd_length, lveq_length_eq, sle_refl, ex4_4_intro/ -qed. - -(* Forward lemmas with free variables inclusion for restricted closures *****) - -lemma lfeq_lfxs_trans: ∀R. lfeq_transitive R → - ∀L1,L,T. L1 ≡[T] L → ∀L2. L ⪤*[R, T] L2 → L1 ⪤*[R, T] L2. -/4 width=16 by lfeq_fsle_comp, lfxs_trans_fsle, lfxs_trans_next/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs.ma deleted file mode 100644 index c9d55b73f..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs.ma +++ /dev/null @@ -1,320 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "ground_2/relocation/rtmap_id.ma". -include "basic_2/notation/relations/relationstar_4.ma". -include "basic_2/syntax/cext2.ma". -include "basic_2/relocation/lexs.ma". -include "basic_2/static/frees.ma". - -(* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****) - -definition lfxs (R) (T): relation lenv ≝ - λL1,L2. ∃∃f. L1 ⊢ 𝐅*⦃T⦄ ≘ f & L1 ⪤*[cext2 R, cfull, f] L2. - -interpretation "generic extension on referred entries (local environment)" - 'RelationStar R T L1 L2 = (lfxs R T L1 L2). - -definition R_confluent2_lfxs: relation4 (relation3 lenv term term) - (relation3 lenv term term) … ≝ - λR1,R2,RP1,RP2. - ∀L0,T0,T1. R1 L0 T0 T1 → ∀T2. R2 L0 T0 T2 → - ∀L1. L0 ⪤*[RP1, T0] L1 → ∀L2. L0 ⪤*[RP2, T0] L2 → - ∃∃T. R2 L1 T1 T & R1 L2 T2 T. - -definition lfxs_confluent: relation … ≝ - λR1,R2. - ∀K1,K,V1. K1 ⪤*[R1, V1] K → ∀V. R1 K1 V1 V → - ∀K2. K ⪤*[R2, V] K2 → K ⪤*[R2, V1] K2. - -definition lfxs_transitive: relation3 ? (relation3 ?? term) … ≝ - λR1,R2,R3. - ∀K1,K,V1. K1 ⪤*[R1, V1] K → - ∀V. R1 K1 V1 V → ∀V2. R2 K V V2 → R3 K1 V1 V2. - -(* Basic inversion lemmas ***************************************************) - -lemma lfxs_inv_atom_sn (R): ∀Y2,T. ⋆ ⪤*[R, T] Y2 → Y2 = ⋆. -#R #Y2 #T * /2 width=4 by lexs_inv_atom1/ -qed-. - -lemma lfxs_inv_atom_dx (R): ∀Y1,T. Y1 ⪤*[R, T] ⋆ → Y1 = ⋆. -#R #I #Y1 * /2 width=4 by lexs_inv_atom2/ -qed-. - -lemma lfxs_inv_sort (R): ∀Y1,Y2,s. Y1 ⪤*[R, ⋆s] Y2 → - ∨∨ Y1 = ⋆ ∧ Y2 = ⋆ - | ∃∃I1,I2,L1,L2. L1 ⪤*[R, ⋆s] L2 & - Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. -#R * [ | #Y1 #I1 ] #Y2 #s * #f #H1 #H2 -[ lapply (lexs_inv_atom1 … H2) -H2 /3 width=1 by or_introl, conj/ -| lapply (frees_inv_sort … H1) -H1 #Hf - elim (isid_inv_gen … Hf) -Hf #g #Hg #H destruct - elim (lexs_inv_push1 … H2) -H2 #I2 #L2 #H12 #_ #H destruct - /5 width=7 by frees_sort, ex3_4_intro, ex2_intro, or_intror/ -] -qed-. - -lemma lfxs_inv_zero (R): ∀Y1,Y2. Y1 ⪤*[R, #0] Y2 → - ∨∨ Y1 = ⋆ ∧ Y2 = ⋆ - | ∃∃I,L1,L2,V1,V2. L1 ⪤*[R, V1] L2 & R L1 V1 V2 & - Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2 - | ∃∃f,I,L1,L2. 𝐈⦃f⦄ & L1 ⪤*[cext2 R, cfull, f] L2 & - Y1 = L1.ⓤ{I} & Y2 = L2.ⓤ{I}. -#R * [ | #Y1 * #I1 [ | #X ] ] #Y2 * #f #H1 #H2 -[ lapply (lexs_inv_atom1 … H2) -H2 /3 width=1 by or3_intro0, conj/ -| elim (frees_inv_unit … H1) -H1 #g #HX #H destruct - elim (lexs_inv_next1 … H2) -H2 #I2 #L2 #HL12 #H #H2 destruct - >(ext2_inv_unit_sn … H) -H /3 width=8 by or3_intro2, ex4_4_intro/ -| elim (frees_inv_pair … H1) -H1 #g #Hg #H destruct - elim (lexs_inv_next1 … H2) -H2 #Z2 #L2 #HL12 #H - elim (ext2_inv_pair_sn … H) -H - /4 width=9 by or3_intro1, ex4_5_intro, ex2_intro/ -] -qed-. - -lemma lfxs_inv_lref (R): ∀Y1,Y2,i. Y1 ⪤*[R, #↑i] Y2 → - ∨∨ Y1 = ⋆ ∧ Y2 = ⋆ - | ∃∃I1,I2,L1,L2. L1 ⪤*[R, #i] L2 & - Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. -#R * [ | #Y1 #I1 ] #Y2 #i * #f #H1 #H2 -[ lapply (lexs_inv_atom1 … H2) -H2 /3 width=1 by or_introl, conj/ -| elim (frees_inv_lref … H1) -H1 #g #Hg #H destruct - elim (lexs_inv_push1 … H2) -H2 - /4 width=7 by ex3_4_intro, ex2_intro, or_intror/ -] -qed-. - -lemma lfxs_inv_gref (R): ∀Y1,Y2,l. Y1 ⪤*[R, §l] Y2 → - ∨∨ Y1 = ⋆ ∧ Y2 = ⋆ - | ∃∃I1,I2,L1,L2. L1 ⪤*[R, §l] L2 & - Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. -#R * [ | #Y1 #I1 ] #Y2 #l * #f #H1 #H2 -[ lapply (lexs_inv_atom1 … H2) -H2 /3 width=1 by or_introl, conj/ -| lapply (frees_inv_gref … H1) -H1 #Hf - elim (isid_inv_gen … Hf) -Hf #g #Hg #H destruct - elim (lexs_inv_push1 … H2) -H2 #I2 #L2 #H12 #_ #H destruct - /5 width=7 by frees_gref, ex3_4_intro, ex2_intro, or_intror/ -] -qed-. - -(* Basic_2A1: uses: llpx_sn_inv_bind llpx_sn_inv_bind_O *) -lemma lfxs_inv_bind (R): ∀p,I,L1,L2,V1,V2,T. L1 ⪤*[R, ⓑ{p,I}V1.T] L2 → R L1 V1 V2 → - ∧∧ L1 ⪤*[R, V1] L2 & L1.ⓑ{I}V1 ⪤*[R, T] L2.ⓑ{I}V2. -#R #p #I #L1 #L2 #V1 #V2 #T * #f #Hf #HL #HV elim (frees_inv_bind … Hf) -Hf -/6 width=6 by sle_lexs_trans, lexs_inv_tl, ext2_pair, sor_inv_sle_dx, sor_inv_sle_sn, ex2_intro, conj/ -qed-. - -(* Basic_2A1: uses: llpx_sn_inv_flat *) -lemma lfxs_inv_flat (R): ∀I,L1,L2,V,T. L1 ⪤*[R, ⓕ{I}V.T] L2 → - ∧∧ L1 ⪤*[R, V] L2 & L1 ⪤*[R, T] L2. -#R #I #L1 #L2 #V #T * #f #Hf #HL elim (frees_inv_flat … Hf) -Hf -/5 width=6 by sle_lexs_trans, sor_inv_sle_dx, sor_inv_sle_sn, ex2_intro, conj/ -qed-. - -(* Advanced inversion lemmas ************************************************) - -lemma lfxs_inv_sort_bind_sn (R): ∀I1,K1,L2,s. K1.ⓘ{I1} ⪤*[R, ⋆s] L2 → - ∃∃I2,K2. K1 ⪤*[R, ⋆s] K2 & L2 = K2.ⓘ{I2}. -#R #I1 #K1 #L2 #s #H elim (lfxs_inv_sort … H) -H * -[ #H destruct -| #Z1 #I2 #Y1 #K2 #Hs #H1 #H2 destruct /2 width=4 by ex2_2_intro/ -] -qed-. - -lemma lfxs_inv_sort_bind_dx (R): ∀I2,K2,L1,s. L1 ⪤*[R, ⋆s] K2.ⓘ{I2} → - ∃∃I1,K1. K1 ⪤*[R, ⋆s] K2 & L1 = K1.ⓘ{I1}. -#R #I2 #K2 #L1 #s #H elim (lfxs_inv_sort … H) -H * -[ #_ #H destruct -| #I1 #Z2 #K1 #Y2 #Hs #H1 #H2 destruct /2 width=4 by ex2_2_intro/ -] -qed-. - -lemma lfxs_inv_zero_pair_sn (R): ∀I,L2,K1,V1. K1.ⓑ{I}V1 ⪤*[R, #0] L2 → - ∃∃K2,V2. K1 ⪤*[R, V1] K2 & R K1 V1 V2 & - L2 = K2.ⓑ{I}V2. -#R #I #L2 #K1 #V1 #H elim (lfxs_inv_zero … H) -H * -[ #H destruct -| #Z #Y1 #K2 #X1 #V2 #HK12 #HV12 #H1 #H2 destruct - /2 width=5 by ex3_2_intro/ -| #f #Z #Y1 #Y2 #_ #_ #H destruct -] -qed-. - -lemma lfxs_inv_zero_pair_dx (R): ∀I,L1,K2,V2. L1 ⪤*[R, #0] K2.ⓑ{I}V2 → - ∃∃K1,V1. K1 ⪤*[R, V1] K2 & R K1 V1 V2 & - L1 = K1.ⓑ{I}V1. -#R #I #L1 #K2 #V2 #H elim (lfxs_inv_zero … H) -H * -[ #_ #H destruct -| #Z #K1 #Y2 #V1 #X2 #HK12 #HV12 #H1 #H2 destruct - /2 width=5 by ex3_2_intro/ -| #f #Z #Y1 #Y2 #_ #_ #_ #H destruct -] -qed-. - -lemma lfxs_inv_zero_unit_sn (R): ∀I,K1,L2. K1.ⓤ{I} ⪤*[R, #0] L2 → - ∃∃f,K2. 𝐈⦃f⦄ & K1 ⪤*[cext2 R, cfull, f] K2 & - L2 = K2.ⓤ{I}. -#R #I #K1 #L2 #H elim (lfxs_inv_zero … H) -H * -[ #H destruct -| #Z #Y1 #Y2 #X1 #X2 #_ #_ #H destruct -| #f #Z #Y1 #K2 #Hf #HK12 #H1 #H2 destruct /2 width=5 by ex3_2_intro/ -] -qed-. - -lemma lfxs_inv_zero_unit_dx (R): ∀I,L1,K2. L1 ⪤*[R, #0] K2.ⓤ{I} → - ∃∃f,K1. 𝐈⦃f⦄ & K1 ⪤*[cext2 R, cfull, f] K2 & - L1 = K1.ⓤ{I}. -#R #I #L1 #K2 #H elim (lfxs_inv_zero … H) -H * -[ #_ #H destruct -| #Z #Y1 #Y2 #X1 #X2 #_ #_ #_ #H destruct -| #f #Z #K1 #Y2 #Hf #HK12 #H1 #H2 destruct /2 width=5 by ex3_2_intro/ -] -qed-. - -lemma lfxs_inv_lref_bind_sn (R): ∀I1,K1,L2,i. K1.ⓘ{I1} ⪤*[R, #↑i] L2 → - ∃∃I2,K2. K1 ⪤*[R, #i] K2 & L2 = K2.ⓘ{I2}. -#R #I1 #K1 #L2 #i #H elim (lfxs_inv_lref … H) -H * -[ #H destruct -| #Z1 #I2 #Y1 #K2 #Hi #H1 #H2 destruct /2 width=4 by ex2_2_intro/ -] -qed-. - -lemma lfxs_inv_lref_bind_dx (R): ∀I2,K2,L1,i. L1 ⪤*[R, #↑i] K2.ⓘ{I2} → - ∃∃I1,K1. K1 ⪤*[R, #i] K2 & L1 = K1.ⓘ{I1}. -#R #I2 #K2 #L1 #i #H elim (lfxs_inv_lref … H) -H * -[ #_ #H destruct -| #I1 #Z2 #K1 #Y2 #Hi #H1 #H2 destruct /2 width=4 by ex2_2_intro/ -] -qed-. - -lemma lfxs_inv_gref_bind_sn (R): ∀I1,K1,L2,l. K1.ⓘ{I1} ⪤*[R, §l] L2 → - ∃∃I2,K2. K1 ⪤*[R, §l] K2 & L2 = K2.ⓘ{I2}. -#R #I1 #K1 #L2 #l #H elim (lfxs_inv_gref … H) -H * -[ #H destruct -| #Z1 #I2 #Y1 #K2 #Hl #H1 #H2 destruct /2 width=4 by ex2_2_intro/ -] -qed-. - -lemma lfxs_inv_gref_bind_dx (R): ∀I2,K2,L1,l. L1 ⪤*[R, §l] K2.ⓘ{I2} → - ∃∃I1,K1. K1 ⪤*[R, §l] K2 & L1 = K1.ⓘ{I1}. -#R #I2 #K2 #L1 #l #H elim (lfxs_inv_gref … H) -H * -[ #_ #H destruct -| #I1 #Z2 #K1 #Y2 #Hl #H1 #H2 destruct /2 width=4 by ex2_2_intro/ -] -qed-. - -(* Basic forward lemmas *****************************************************) - -lemma lfxs_fwd_zero_pair (R): ∀I,K1,K2,V1,V2. - K1.ⓑ{I}V1 ⪤*[R, #0] K2.ⓑ{I}V2 → K1 ⪤*[R, V1] K2. -#R #I #K1 #K2 #V1 #V2 #H -elim (lfxs_inv_zero_pair_sn … H) -H #Y #X #HK12 #_ #H destruct // -qed-. - -(* Basic_2A1: uses: llpx_sn_fwd_pair_sn llpx_sn_fwd_bind_sn llpx_sn_fwd_flat_sn *) -lemma lfxs_fwd_pair_sn (R): ∀I,L1,L2,V,T. L1 ⪤*[R, ②{I}V.T] L2 → L1 ⪤*[R, V] L2. -#R * [ #p ] #I #L1 #L2 #V #T * #f #Hf #HL -[ elim (frees_inv_bind … Hf) | elim (frees_inv_flat … Hf) ] -Hf -/4 width=6 by sle_lexs_trans, sor_inv_sle_sn, ex2_intro/ -qed-. - -(* Basic_2A1: uses: llpx_sn_fwd_bind_dx llpx_sn_fwd_bind_O_dx *) -lemma lfxs_fwd_bind_dx (R): ∀p,I,L1,L2,V1,V2,T. L1 ⪤*[R, ⓑ{p,I}V1.T] L2 → - R L1 V1 V2 → L1.ⓑ{I}V1 ⪤*[R, T] L2.ⓑ{I}V2. -#R #p #I #L1 #L2 #V1 #V2 #T #H #HV elim (lfxs_inv_bind … H HV) -H -HV // -qed-. - -(* Basic_2A1: uses: llpx_sn_fwd_flat_dx *) -lemma lfxs_fwd_flat_dx (R): ∀I,L1,L2,V,T. L1 ⪤*[R, ⓕ{I}V.T] L2 → L1 ⪤*[R, T] L2. -#R #I #L1 #L2 #V #T #H elim (lfxs_inv_flat … H) -H // -qed-. - -lemma lfxs_fwd_dx (R): ∀I2,L1,K2,T. L1 ⪤*[R, T] K2.ⓘ{I2} → - ∃∃I1,K1. L1 = K1.ⓘ{I1}. -#R #I2 #L1 #K2 #T * #f elim (pn_split f) * #g #Hg #_ #Hf destruct -[ elim (lexs_inv_push2 … Hf) | elim (lexs_inv_next2 … Hf) ] -Hf #I1 #K1 #_ #_ #H destruct -/2 width=3 by ex1_2_intro/ -qed-. - -(* Basic properties *********************************************************) - -lemma lfxs_atom (R): ∀I. ⋆ ⪤*[R, ⓪{I}] ⋆. -#R * /3 width=3 by frees_sort, frees_atom, frees_gref, lexs_atom, ex2_intro/ -qed. - -lemma lfxs_sort (R): ∀I1,I2,L1,L2,s. - L1 ⪤*[R, ⋆s] L2 → L1.ⓘ{I1} ⪤*[R, ⋆s] L2.ⓘ{I2}. -#R #I1 #I2 #L1 #L2 #s * #f #Hf #H12 -lapply (frees_inv_sort … Hf) -Hf -/4 width=3 by frees_sort, lexs_push, isid_push, ex2_intro/ -qed. - -lemma lfxs_pair (R): ∀I,L1,L2,V1,V2. L1 ⪤*[R, V1] L2 → - R L1 V1 V2 → L1.ⓑ{I}V1 ⪤*[R, #0] L2.ⓑ{I}V2. -#R #I1 #I2 #L1 #L2 #V1 * -/4 width=3 by ext2_pair, frees_pair, lexs_next, ex2_intro/ -qed. - -lemma lfxs_unit (R): ∀f,I,L1,L2. 𝐈⦃f⦄ → L1 ⪤*[cext2 R, cfull, f] L2 → - L1.ⓤ{I} ⪤*[R, #0] L2.ⓤ{I}. -/4 width=3 by frees_unit, lexs_next, ext2_unit, ex2_intro/ qed. - -lemma lfxs_lref (R): ∀I1,I2,L1,L2,i. - L1 ⪤*[R, #i] L2 → L1.ⓘ{I1} ⪤*[R, #↑i] L2.ⓘ{I2}. -#R #I1 #I2 #L1 #L2 #i * /3 width=3 by lexs_push, frees_lref, ex2_intro/ -qed. - -lemma lfxs_gref (R): ∀I1,I2,L1,L2,l. - L1 ⪤*[R, §l] L2 → L1.ⓘ{I1} ⪤*[R, §l] L2.ⓘ{I2}. -#R #I1 #I2 #L1 #L2 #l * #f #Hf #H12 -lapply (frees_inv_gref … Hf) -Hf -/4 width=3 by frees_gref, lexs_push, isid_push, ex2_intro/ -qed. - -lemma lfxs_bind_repl_dx (R): ∀I,I1,L1,L2,T. - L1.ⓘ{I} ⪤*[R, T] L2.ⓘ{I1} → - ∀I2. cext2 R L1 I I2 → - L1.ⓘ{I} ⪤*[R, T] L2.ⓘ{I2}. -#R #I #I1 #L1 #L2 #T * #f #Hf #HL12 #I2 #HR -/3 width=5 by lexs_pair_repl, ex2_intro/ -qed-. - -(* Basic_2A1: uses: llpx_sn_co *) -lemma lfxs_co (R1) (R2): (∀L,T1,T2. R1 L T1 T2 → R2 L T1 T2) → - ∀L1,L2,T. L1 ⪤*[R1, T] L2 → L1 ⪤*[R2, T] L2. -#R1 #R2 #HR #L1 #L2 #T * /5 width=7 by lexs_co, cext2_co, ex2_intro/ -qed-. - -lemma lfxs_isid (R1) (R2): ∀L1,L2,T1,T2. - (∀f. L1 ⊢ 𝐅*⦃T1⦄ ≘ f → 𝐈⦃f⦄) → - (∀f. 𝐈⦃f⦄ → L1 ⊢ 𝐅*⦃T2⦄ ≘ f) → - L1 ⪤*[R1, T1] L2 → L1 ⪤*[R2, T2] L2. -#R1 #R2 #L1 #L2 #T1 #T2 #H1 #H2 * -/4 width=7 by lexs_co_isid, ex2_intro/ -qed-. - -lemma lfxs_unit_sn (R1) (R2): - ∀I,K1,L2. K1.ⓤ{I} ⪤*[R1, #0] L2 → K1.ⓤ{I} ⪤*[R2, #0] L2. -#R1 #R2 #I #K1 #L2 #H -elim (lfxs_inv_zero_unit_sn … H) -H #f #K2 #Hf #HK12 #H destruct -/3 width=7 by lfxs_unit, lexs_co_isid/ -qed-. - -(* Basic_2A1: removed theorems 9: - llpx_sn_skip llpx_sn_lref llpx_sn_free - llpx_sn_fwd_lref - llpx_sn_Y llpx_sn_ge_up llpx_sn_ge - llpx_sn_fwd_drop_sn llpx_sn_fwd_drop_dx -*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_drops.ma deleted file mode 100644 index 76adfa38b..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_drops.ma +++ /dev/null @@ -1,124 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/relocation/drops_cext2.ma". -include "basic_2/relocation/drops_lexs.ma". -include "basic_2/static/frees_drops.ma". -include "basic_2/static/lfxs.ma". - -(* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****) - -definition f_dedropable_sn: predicate (relation3 lenv term term) ≝ - λR. ∀b,f,L1,K1. ⬇*[b, f] L1 ≘ K1 → - ∀K2,T. K1 ⪤*[R, T] K2 → ∀U. ⬆*[f] T ≘ U → - ∃∃L2. L1 ⪤*[R, U] L2 & ⬇*[b, f] L2 ≘ K2 & L1 ≡[f] L2. - -definition f_dropable_sn: predicate (relation3 lenv term term) ≝ - λR. ∀b,f,L1,K1. ⬇*[b, f] L1 ≘ K1 → 𝐔⦃f⦄ → - ∀L2,U. L1 ⪤*[R, U] L2 → ∀T. ⬆*[f] T ≘ U → - ∃∃K2. K1 ⪤*[R, T] K2 & ⬇*[b, f] L2 ≘ K2. - -definition f_dropable_dx: predicate (relation3 lenv term term) ≝ - λR. ∀L1,L2,U. L1 ⪤*[R, U] L2 → - ∀b,f,K2. ⬇*[b, f] L2 ≘ K2 → 𝐔⦃f⦄ → ∀T. ⬆*[f] T ≘ U → - ∃∃K1. ⬇*[b, f] L1 ≘ K1 & K1 ⪤*[R, T] K2. - -definition f_transitive_next: relation3 … ≝ λR1,R2,R3. - ∀f,L,T. L ⊢ 𝐅*⦃T⦄ ≘ f → - ∀g,I,K,n. ⬇*[n] L ≘ K.ⓘ{I} → ↑g = ⫱*[n] f → - lexs_transitive (cext2 R1) (cext2 R2) (cext2 R3) (cext2 R1) cfull g K I. - -(* Properties with generic slicing for local environments *******************) - -lemma lfxs_liftable_dedropable_sn: ∀R. (∀L. reflexive ? (R L)) → - d_liftable2_sn … lifts R → f_dedropable_sn R. -#R #H1R #H2R #b #f #L1 #K1 #HLK1 #K2 #T * #f1 #Hf1 #HK12 #U #HTU -elim (frees_total L1 U) #f2 #Hf2 -lapply (frees_fwd_coafter … Hf2 … HLK1 … HTU … Hf1) -HTU #Hf -elim (lexs_liftable_co_dedropable_sn … HLK1 … HK12 … Hf) -f1 -K1 -/3 width=6 by cext2_d_liftable2_sn, cfull_lift_sn, ext2_refl, ex3_intro, ex2_intro/ -qed-. - -lemma lfxs_trans_next: ∀R1,R2,R3. lfxs_transitive R1 R2 R3 → f_transitive_next R1 R2 R3. -#R1 #R2 #R3 #HR #f #L1 #T #Hf #g #I1 #K1 #n #HLK #Hgf #I #H -generalize in match HLK; -HLK elim H -I1 -I -[ #I #_ #L2 #_ #I2 #H - lapply (ext2_inv_unit_sn … H) -H #H destruct - /2 width=1 by ext2_unit/ -| #I #V1 #V #HV1 #HLK1 #L2 #HL12 #I2 #H - elim (ext2_inv_pair_sn … H) -H #V2 #HV2 #H destruct - elim (frees_inv_drops_next … Hf … HLK1 … Hgf) -f -HLK1 #f #Hf #Hfg - /5 width=5 by ext2_pair, sle_lexs_trans, ex2_intro/ -] -qed. - -(* Inversion lemmas with generic slicing for local environments *************) - -(* Basic_2A1: uses: llpx_sn_inv_lift_le llpx_sn_inv_lift_be llpx_sn_inv_lift_ge *) -(* Basic_2A1: was: llpx_sn_drop_conf_O *) -lemma lfxs_dropable_sn: ∀R. f_dropable_sn R. -#R #b #f #L1 #K1 #HLK1 #H1f #L2 #U * #f2 #Hf2 #HL12 #T #HTU -elim (frees_total K1 T) #f1 #Hf1 -lapply (frees_fwd_coafter … Hf2 … HLK1 … HTU … Hf1) -HTU #H2f -elim (lexs_co_dropable_sn … HLK1 … HL12 … H2f) -f2 -L1 -/3 width=3 by ex2_intro/ -qed-. - -(* Basic_2A1: was: llpx_sn_drop_trans_O *) -(* Note: the proof might be simplified *) -lemma lfxs_dropable_dx: ∀R. f_dropable_dx R. -#R #L1 #L2 #U * #f2 #Hf2 #HL12 #b #f #K2 #HLK2 #H1f #T #HTU -elim (drops_isuni_ex … H1f L1) #K1 #HLK1 -elim (frees_total K1 T) #f1 #Hf1 -lapply (frees_fwd_coafter … Hf2 … HLK1 … HTU … Hf1) -K1 #H2f -elim (lexs_co_dropable_dx … HL12 … HLK2 … H2f) -L2 -/4 width=9 by frees_inv_lifts, ex2_intro/ -qed-. - -(* Basic_2A1: uses: llpx_sn_inv_lift_O *) -lemma lfxs_inv_lifts_bi: ∀R,L1,L2,U. L1 ⪤*[R, U] L2 → ∀b,f. 𝐔⦃f⦄ → - ∀K1,K2. ⬇*[b, f] L1 ≘ K1 → ⬇*[b, f] L2 ≘ K2 → - ∀T. ⬆*[f] T ≘ U → K1 ⪤*[R, T] K2. -#R #L1 #L2 #U #HL12 #b #f #Hf #K1 #K2 #HLK1 #HLK2 #T #HTU -elim (lfxs_dropable_sn … HLK1 … HL12 … HTU) -L1 -U // #Y #HK12 #HY -lapply (drops_mono … HY … HLK2) -b -f -L2 #H destruct // -qed-. - -lemma lfxs_inv_lref_pair_sn: ∀R,L1,L2,i. L1 ⪤*[R, #i] L2 → ∀I,K1,V1. ⬇*[i] L1 ≘ K1.ⓑ{I}V1 → - ∃∃K2,V2. ⬇*[i] L2 ≘ K2.ⓑ{I}V2 & K1 ⪤*[R, V1] K2 & R K1 V1 V2. -#R #L1 #L2 #i #HL12 #I #K1 #V1 #HLK1 elim (lfxs_dropable_sn … HLK1 … HL12 (#0)) -HLK1 -HL12 // -#Y #HY #HLK2 elim (lfxs_inv_zero_pair_sn … HY) -HY -#K2 #V2 #HK12 #HV12 #H destruct /2 width=5 by ex3_2_intro/ -qed-. - -lemma lfxs_inv_lref_pair_dx: ∀R,L1,L2,i. L1 ⪤*[R, #i] L2 → ∀I,K2,V2. ⬇*[i] L2 ≘ K2.ⓑ{I}V2 → - ∃∃K1,V1. ⬇*[i] L1 ≘ K1.ⓑ{I}V1 & K1 ⪤*[R, V1] K2 & R K1 V1 V2. -#R #L1 #L2 #i #HL12 #I #K2 #V2 #HLK2 elim (lfxs_dropable_dx … HL12 … HLK2 … (#0)) -HLK2 -HL12 // -#Y #HLK1 #HY elim (lfxs_inv_zero_pair_dx … HY) -HY -#K1 #V1 #HK12 #HV12 #H destruct /2 width=5 by ex3_2_intro/ -qed-. - -lemma lfxs_inv_lref_unit_sn: ∀R,L1,L2,i. L1 ⪤*[R, #i] L2 → ∀I,K1. ⬇*[i] L1 ≘ K1.ⓤ{I} → - ∃∃f,K2. ⬇*[i] L2 ≘ K2.ⓤ{I} & K1 ⪤*[cext2 R, cfull, f] K2 & 𝐈⦃f⦄. -#R #L1 #L2 #i #HL12 #I #K1 #HLK1 elim (lfxs_dropable_sn … HLK1 … HL12 (#0)) -HLK1 -HL12 // -#Y #HY #HLK2 elim (lfxs_inv_zero_unit_sn … HY) -HY -#f #K2 #Hf #HK12 #H destruct /2 width=5 by ex3_2_intro/ -qed-. - -lemma lfxs_inv_lref_unit_dx: ∀R,L1,L2,i. L1 ⪤*[R, #i] L2 → ∀I,K2. ⬇*[i] L2 ≘ K2.ⓤ{I} → - ∃∃f,K1. ⬇*[i] L1 ≘ K1.ⓤ{I} & K1 ⪤*[cext2 R, cfull, f] K2 & 𝐈⦃f⦄. -#R #L1 #L2 #i #HL12 #I #K2 #HLK2 elim (lfxs_dropable_dx … HL12 … HLK2 … (#0)) -HLK2 -HL12 // -#Y #HLK1 #HY elim (lfxs_inv_zero_unit_dx … HY) -HY -#f #K2 #Hf #HK12 #H destruct /2 width=5 by ex3_2_intro/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_fqup.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_fqup.ma deleted file mode 100644 index 6bf9a2697..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_fqup.ma +++ /dev/null @@ -1,49 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/frees_fqup.ma". -include "basic_2/static/lfxs.ma". - -(* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****) - -(* Advanced properties ******************************************************) - -(* Basic_2A1: uses: llpx_sn_refl *) -lemma lfxs_refl: ∀R. (∀L. reflexive … (R L)) → ∀L,T. L ⪤*[R, T] L. -#R #HR #L #T elim (frees_total L T) -/4 width=3 by lexs_refl, ext2_refl, ex2_intro/ -qed. - -lemma lfxs_pair_refl: ∀R. (∀L. reflexive … (R L)) → - ∀L,V1,V2. R L V1 V2 → ∀I,T. L.ⓑ{I}V1 ⪤*[R, T] L.ⓑ{I}V2. -#R #HR #L #V1 #V2 #HV12 #I #T -elim (frees_total (L.ⓑ{I}V1) T) #f #Hf -elim (pn_split f) * #g #H destruct -/5 width=3 by lexs_refl, lexs_next, lexs_push, ext2_refl, ext2_pair, ex2_intro/ -qed. - -(* Advanced inversion lemmas ************************************************) - -lemma lfxs_inv_bind_void: ∀R,p,I,L1,L2,V,T. L1 ⪤*[R, ⓑ{p,I}V.T] L2 → - L1 ⪤*[R, V] L2 ∧ L1.ⓧ ⪤*[R, T] L2.ⓧ. -#R #p #I #L1 #L2 #V #T * #f #Hf #HL elim (frees_inv_bind_void … Hf) -Hf -/6 width=6 by sle_lexs_trans, lexs_inv_tl, sor_inv_sle_dx, sor_inv_sle_sn, ex2_intro, conj/ -qed-. - -(* Advanced forward lemmas **************************************************) - -lemma lfxs_fwd_bind_dx_void: ∀R,p,I,L1,L2,V,T. L1 ⪤*[R, ⓑ{p,I}V.T] L2 → - L1.ⓧ ⪤*[R, T] L2.ⓧ. -#R #p #I #L1 #L2 #V #T #H elim (lfxs_inv_bind_void … H) -H // -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_fsle.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_fsle.ma deleted file mode 100644 index 514e63def..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_fsle.ma +++ /dev/null @@ -1,176 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/relocation/lexs_length.ma". -include "basic_2/static/fsle_fsle.ma". -include "basic_2/static/lfxs_drops.ma". -include "basic_2/static/lfxs_lfxs.ma". - -(* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****) - -definition R_fsge_compatible: predicate (relation3 …) ≝ λRN. - ∀L,T1,T2. RN L T1 T2 → ⦃L, T2⦄ ⊆ ⦃L, T1⦄. - -definition lfxs_fsge_compatible: predicate (relation3 …) ≝ λRN. - ∀L1,L2,T. L1 ⪤*[RN, T] L2 → ⦃L2, T⦄ ⊆ ⦃L1, T⦄. - -definition lfxs_fsle_compatible: predicate (relation3 …) ≝ λRN. - ∀L1,L2,T. L1 ⪤*[RN, T] L2 → ⦃L1, T⦄ ⊆ ⦃L2, T⦄. - -(* Basic inversions with free variables inclusion for restricted closures ***) - -lemma frees_lexs_conf: ∀R. lfxs_fsge_compatible R → - ∀L1,T,f1. L1 ⊢ 𝐅*⦃T⦄ ≘ f1 → - ∀L2. L1 ⪤*[cext2 R, cfull, f1] L2 → - ∃∃f2. L2 ⊢ 𝐅*⦃T⦄ ≘ f2 & f2 ⊆ f1. -#R #HR #L1 #T #f1 #Hf1 #L2 #H1L -lapply (HR L1 L2 T ?) /2 width=3 by ex2_intro/ #H2L -@(fsle_frees_trans_eq … H2L … Hf1) /3 width=4 by lexs_fwd_length, sym_eq/ -qed-. - -(* Properties with free variables inclusion for restricted closures *********) - -(* Note: we just need lveq_inv_refl: ∀L,n1,n2. L ≋ⓧ*[n1, n2] L → ∧∧ 0 = n1 & 0 = n2 *) -lemma fsge_lfxs_trans: ∀R,L1,T1,T2. ⦃L1, T1⦄ ⊆ ⦃L1, T2⦄ → - ∀L2. L1 ⪤*[R, T2] L2 → L1 ⪤*[R, T1] L2. -#R #L1 #T1 #T2 * #n1 #n2 #f1 #f2 #Hf1 #Hf2 #Hn #Hf #L2 #HL12 -elim (lveq_inj_length … Hn ?) // #H1 #H2 destruct -/4 width=5 by lfxs_inv_frees, sle_lexs_trans, ex2_intro/ -qed-. - -lemma lfxs_sym: ∀R. lfxs_fsge_compatible R → - (∀L1,L2,T1,T2. R L1 T1 T2 → R L2 T2 T1) → - ∀T. symmetric … (lfxs R T). -#R #H1R #H2R #T #L1 #L2 -* #f1 #Hf1 #HL12 -elim (frees_lexs_conf … Hf1 … HL12) -Hf1 // -/5 width=5 by sle_lexs_trans, lexs_sym, cext2_sym, ex2_intro/ -qed-. - -lemma lfxs_pair_sn_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) → - lfxs_fsge_compatible R1 → - ∀L1,L2,V. L1 ⪤*[R1, V] L2 → ∀I,T. - ∃∃L. L1 ⪤*[R1, ②{I}V.T] L & L ⪤*[R2, V] L2. -#R1 #R2 #HR1 #HR2 #HR #L1 #L2 #V * #f #Hf #HL12 * [ #p ] #I #T -[ elim (frees_total L1 (ⓑ{p,I}V.T)) #g #Hg - elim (frees_inv_bind … Hg) #y1 #y2 #H #_ #Hy -| elim (frees_total L1 (ⓕ{I}V.T)) #g #Hg - elim (frees_inv_flat … Hg) #y1 #y2 #H #_ #Hy -] -lapply(frees_mono … H … Hf) -H #H1 -lapply (sor_eq_repl_back1 … Hy … H1) -y1 #Hy -lapply (sor_inv_sle_sn … Hy) -y2 #Hfg -elim (lexs_sle_split (cext2 R1) (cext2 R2) … HL12 … Hfg) -HL12 /2 width=1 by ext2_refl/ #L #HL1 #HL2 -lapply (sle_lexs_trans … HL1 … Hfg) // #H -elim (frees_lexs_conf … Hf … H) -Hf -H -/4 width=7 by sle_lexs_trans, ex2_intro/ -qed-. - -lemma lfxs_flat_dx_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) → - lfxs_fsge_compatible R1 → - ∀L1,L2,T. L1 ⪤*[R1, T] L2 → ∀I,V. - ∃∃L. L1 ⪤*[R1, ⓕ{I}V.T] L & L ⪤*[R2, T] L2. -#R1 #R2 #HR1 #HR2 #HR #L1 #L2 #T * #f #Hf #HL12 #I #V -elim (frees_total L1 (ⓕ{I}V.T)) #g #Hg -elim (frees_inv_flat … Hg) #y1 #y2 #_ #H #Hy -lapply(frees_mono … H … Hf) -H #H2 -lapply (sor_eq_repl_back2 … Hy … H2) -y2 #Hy -lapply (sor_inv_sle_dx … Hy) -y1 #Hfg -elim (lexs_sle_split (cext2 R1) (cext2 R2) … HL12 … Hfg) -HL12 /2 width=1 by ext2_refl/ #L #HL1 #HL2 -lapply (sle_lexs_trans … HL1 … Hfg) // #H -elim (frees_lexs_conf … Hf … H) -Hf -H -/4 width=7 by sle_lexs_trans, ex2_intro/ -qed-. - -lemma lfxs_bind_dx_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) → - lfxs_fsge_compatible R1 → - ∀I,L1,L2,V1,T. L1.ⓑ{I}V1 ⪤*[R1, T] L2 → ∀p. - ∃∃L,V. L1 ⪤*[R1, ⓑ{p,I}V1.T] L & L.ⓑ{I}V ⪤*[R2, T] L2 & R1 L1 V1 V. -#R1 #R2 #HR1 #HR2 #HR #I #L1 #L2 #V1 #T * #f #Hf #HL12 #p -elim (frees_total L1 (ⓑ{p,I}V1.T)) #g #Hg -elim (frees_inv_bind … Hg) #y1 #y2 #_ #H #Hy -lapply(frees_mono … H … Hf) -H #H2 -lapply (tl_eq_repl … H2) -H2 #H2 -lapply (sor_eq_repl_back2 … Hy … H2) -y2 #Hy -lapply (sor_inv_sle_dx … Hy) -y1 #Hfg -lapply (sle_inv_tl_sn … Hfg) -Hfg #Hfg -elim (lexs_sle_split (cext2 R1) (cext2 R2) … HL12 … Hfg) -HL12 /2 width=1 by ext2_refl/ #Y #H #HL2 -lapply (sle_lexs_trans … H … Hfg) // #H0 -elim (lexs_inv_next1 … H) -H #Z #L #HL1 #H -elim (ext2_inv_pair_sn … H) -H #V #HV #H1 #H2 destruct -elim (frees_lexs_conf … Hf … H0) -Hf -H0 -/4 width=7 by sle_lexs_trans, ex3_2_intro, ex2_intro/ -qed-. - -lemma lfxs_bind_dx_split_void: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) → - lfxs_fsge_compatible R1 → - ∀L1,L2,T. L1.ⓧ ⪤*[R1, T] L2 → ∀p,I,V. - ∃∃L. L1 ⪤*[R1, ⓑ{p,I}V.T] L & L.ⓧ ⪤*[R2, T] L2. -#R1 #R2 #HR1 #HR2 #HR #L1 #L2 #T * #f #Hf #HL12 #p #I #V -elim (frees_total L1 (ⓑ{p,I}V.T)) #g #Hg -elim (frees_inv_bind_void … Hg) #y1 #y2 #_ #H #Hy -lapply(frees_mono … H … Hf) -H #H2 -lapply (tl_eq_repl … H2) -H2 #H2 -lapply (sor_eq_repl_back2 … Hy … H2) -y2 #Hy -lapply (sor_inv_sle_dx … Hy) -y1 #Hfg -lapply (sle_inv_tl_sn … Hfg) -Hfg #Hfg -elim (lexs_sle_split (cext2 R1) (cext2 R2) … HL12 … Hfg) -HL12 /2 width=1 by ext2_refl/ #Y #H #HL2 -lapply (sle_lexs_trans … H … Hfg) // #H0 -elim (lexs_inv_next1 … H) -H #Z #L #HL1 #H -elim (ext2_inv_unit_sn … H) -H #H destruct -elim (frees_lexs_conf … Hf … H0) -Hf -H0 -/4 width=7 by sle_lexs_trans, ex2_intro/ (* note: 2 ex2_intro *) -qed-. - -(* Main properties with free variables inclusion for restricted closures ****) - -theorem lfxs_conf: ∀R1,R2. - lfxs_fsge_compatible R1 → - lfxs_fsge_compatible R2 → - R_confluent2_lfxs R1 R2 R1 R2 → - ∀T. confluent2 … (lfxs R1 T) (lfxs R2 T). -#R1 #R2 #HR1 #HR2 #HR12 #T #L0 #L1 * #f1 #Hf1 #HL01 #L2 * #f #Hf #HL02 -lapply (frees_mono … Hf1 … Hf) -Hf1 #Hf12 -lapply (lexs_eq_repl_back … HL01 … Hf12) -f1 #HL01 -elim (lexs_conf … HL01 … HL02) /2 width=3 by ex2_intro/ [ | -HL01 -HL02 ] -[ #L #HL1 #HL2 - elim (frees_lexs_conf … Hf … HL01) // -HR1 -HL01 #f1 #Hf1 #H1 - elim (frees_lexs_conf … Hf … HL02) // -HR2 -HL02 #f2 #Hf2 #H2 - lapply (sle_lexs_trans … HL1 … H1) // -HL1 -H1 #HL1 - lapply (sle_lexs_trans … HL2 … H2) // -HL2 -H2 #HL2 - /3 width=5 by ex2_intro/ -| #g * #I0 [2: #V0 ] #K0 #n #HLK0 #Hgf #Z1 #H1 #Z2 #H2 #K1 #HK01 #K2 #HK02 - [ elim (ext2_inv_pair_sn … H1) -H1 #V1 #HV01 #H destruct - elim (ext2_inv_pair_sn … H2) -H2 #V2 #HV02 #H destruct - elim (frees_inv_drops_next … Hf … HLK0 … Hgf) -Hf -HLK0 -Hgf #g0 #Hg0 #H0 - lapply (sle_lexs_trans … HK01 … H0) // -HK01 #HK01 - lapply (sle_lexs_trans … HK02 … H0) // -HK02 #HK02 - elim (HR12 … HV01 … HV02 K1 … K2) /3 width=3 by ext2_pair, ex2_intro/ - | lapply (ext2_inv_unit_sn … H1) -H1 #H destruct - lapply (ext2_inv_unit_sn … H2) -H2 #H destruct - /3 width=3 by ext2_unit, ex2_intro/ - ] -] -qed-. - -theorem lfxs_trans_fsle: ∀R1,R2,R3. - lfxs_fsle_compatible R1 → f_transitive_next R1 R2 R3 → - ∀L1,L,T. L1 ⪤*[R1, T] L → - ∀L2. L ⪤*[R2, T] L2 → L1 ⪤*[R3, T] L2. -#R1 #R2 #R3 #H1R #H2R #L1 #L #T #H -lapply (H1R … H) -H1R #H0 -cases H -H #f1 #Hf1 #HL1 #L2 * #f2 #Hf2 #HL2 -lapply (fsle_inv_frees_eq … H0 … Hf1 … Hf2) -H0 -Hf2 -/4 width=14 by lexs_trans_gen, lexs_fwd_length, sle_lexs_trans, ex2_intro/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_length.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_length.ma deleted file mode 100644 index bba7867c7..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_length.ma +++ /dev/null @@ -1,72 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/relocation/lexs_length.ma". -include "basic_2/static/lfxs_drops.ma". - -(* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****) - -(* Forward lemmas with length for local environments ************************) - -(* Basic_2A1: uses: llpx_sn_fwd_length *) -lemma lfxs_fwd_length (R): ∀L1,L2,T. L1 ⪤*[R, T] L2 → |L1| = |L2|. -#R #L1 #L2 #T * /2 width=4 by lexs_fwd_length/ -qed-. - -(* Properties with length for local environments ****************************) - -(* Basic_2A1: uses: llpx_sn_sort *) -lemma lfxs_sort_length (R): ∀L1,L2. |L1| = |L2| → ∀s. L1 ⪤*[R, ⋆s] L2. -#R #L1 elim L1 -L1 -[ #Y #H #s >(length_inv_zero_sn … H) -H // -| #K1 #I1 #IH #Y #H #s - elim (length_inv_succ_sn … H) -H #I2 #K2 #HK12 #H destruct - /3 width=1 by lfxs_sort/ -] -qed. - -(* Basic_2A1: uses: llpx_sn_gref *) -lemma lfxs_gref_length (R): ∀L1,L2. |L1| = |L2| → ∀l. L1 ⪤*[R, §l] L2. -#R #L1 elim L1 -L1 -[ #Y #H #s >(length_inv_zero_sn … H) -H // -| #K1 #I1 #IH #Y #H #s - elim (length_inv_succ_sn … H) -H #I2 #K2 #HK12 #H destruct - /3 width=1 by lfxs_gref/ -] -qed. - -lemma lfxs_unit_length (R): ∀L1,L2. |L1| = |L2| → ∀I. L1.ⓤ{I} ⪤*[R, #0] L2.ⓤ{I}. -/3 width=3 by lfxs_unit, lexs_length_isid/ qed. - -(* Basic_2A1: uses: llpx_sn_lift_le llpx_sn_lift_ge *) -lemma lfxs_lifts_bi (R): d_liftable2_sn … lifts R → - ∀L1,L2. |L1| = |L2| → ∀K1,K2,T. K1 ⪤*[R, T] K2 → - ∀b,f. ⬇*[b, f] L1 ≘ K1 → ⬇*[b, f] L2 ≘ K2 → - ∀U. ⬆*[f] T ≘ U → L1 ⪤*[R, U] L2. -#R #HR #L1 #L2 #HL12 #K1 #K2 #T * #f1 #Hf1 #HK12 #b #f #HLK1 #HLK2 #U #HTU -elim (frees_total L1 U) #f2 #Hf2 -lapply (frees_fwd_coafter … Hf2 … HLK1 … HTU … Hf1) -HTU #Hf -/4 width=12 by lexs_length_cfull, lexs_liftable_co_dedropable_bi, cext2_d_liftable2_sn, cfull_lift_sn, ex2_intro/ -qed-. - -(* Inversion lemmas with length for local environment ***********************) - -lemma lfxs_inv_zero_length (R): ∀Y1,Y2. Y1 ⪤*[R, #0] Y2 → - ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆ - | ∃∃I,L1,L2,V1,V2. L1 ⪤*[R, V1] L2 & R L1 V1 V2 & - Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2 - | ∃∃I,L1,L2. |L1| = |L2| & Y1 = L1.ⓤ{I} & Y2 = L2.ⓤ{I}. -#R #Y1 #Y2 #H elim (lfxs_inv_zero … H) -H * -/4 width=9 by lexs_fwd_length, ex4_5_intro, ex3_3_intro, or3_intro2, or3_intro1, or3_intro0, conj/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lex.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lex.ma deleted file mode 100644 index 01f578f4e..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lex.ma +++ /dev/null @@ -1,41 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/relocation/lex.ma". -include "basic_2/static/lfxs_fsle.ma". -include "basic_2/static/lfeq.ma". - -(* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****) - -(* Properties with generic extension of a context-sensitive relation ********) - -lemma lfxs_lex: ∀R,L1,L2. L1 ⪤[R] L2 → ∀T. L1 ⪤*[R, T] L2. -#R #L1 #L2 * #f #Hf #HL12 #T -elim (frees_total L1 T) #g #Hg -/4 width=5 by lexs_sdj, sdj_isid_sn, ex2_intro/ -qed. - -(* Inversion lemmas with generic extension of a context sensitive relation **) - -lemma lfxs_inv_lex_lfeq: ∀R. c_reflexive … R → - lfxs_fsge_compatible R → - ∀L1,L2,T. L1 ⪤*[R, T] L2 → - ∃∃L. L1 ⪤[R] L & L ≡[T] L2. -#R #H1R #H2R #L1 #L2 #T * #f1 #Hf1 #HL -elim (lexs_sdj_split … ceq_ext … HL 𝐈𝐝 ?) -HL -[ #L0 #HL10 #HL02 |*: /2 width=1 by ext2_refl, sdj_isid_dx/ ] -H1R -lapply (lexs_sdj … HL10 f1 ?) /2 width=1 by sdj_isid_sn/ #H -elim (frees_lexs_conf … Hf1 … H) // -H2R -H #f0 #Hf0 #Hf01 -/4 width=7 by sle_lexs_trans, (* 2x *) ex2_intro/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lfxs.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lfxs.ma deleted file mode 100644 index 89b84566b..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lfxs.ma +++ /dev/null @@ -1,89 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/relocation/lexs_lexs.ma". -include "basic_2/static/frees_fqup.ma". -include "basic_2/static/lfxs.ma". - -(* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****) - -(* Advanced inversion lemmas ************************************************) - -lemma lfxs_inv_frees: ∀R,L1,L2,T. L1 ⪤*[R, T] L2 → - ∀f. L1 ⊢ 𝐅*⦃T⦄ ≘ f → L1 ⪤*[cext2 R, cfull, f] L2. -#R #L1 #L2 #T * /3 width=6 by frees_mono, lexs_eq_repl_back/ -qed-. - -(* Advanced properties ******************************************************) - -(* Basic_2A1: uses: llpx_sn_dec *) -lemma lfxs_dec: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) → - ∀L1,L2,T. Decidable (L1 ⪤*[R, T] L2). -#R #HR #L1 #L2 #T -elim (frees_total L1 T) #f #Hf -elim (lexs_dec (cext2 R) cfull … L1 L2 f) -/4 width=3 by lfxs_inv_frees, cfull_dec, ext2_dec, ex2_intro, or_intror, or_introl/ -qed-. - -(* Main properties **********************************************************) - -(* Basic_2A1: uses: llpx_sn_bind llpx_sn_bind_O *) -theorem lfxs_bind: ∀R,p,I,L1,L2,V1,V2,T. - L1 ⪤*[R, V1] L2 → L1.ⓑ{I}V1 ⪤*[R, T] L2.ⓑ{I}V2 → - L1 ⪤*[R, ⓑ{p,I}V1.T] L2. -#R #p #I #L1 #L2 #V1 #V2 #T * #f1 #HV #Hf1 * #f2 #HT #Hf2 -lapply (lexs_fwd_bind … Hf2) -Hf2 #Hf2 elim (sor_isfin_ex f1 (⫱f2)) -/3 width=7 by frees_fwd_isfin, frees_bind, lexs_join, isfin_tl, ex2_intro/ -qed. - -(* Basic_2A1: llpx_sn_flat *) -theorem lfxs_flat: ∀R,I,L1,L2,V,T. - L1 ⪤*[R, V] L2 → L1 ⪤*[R, T] L2 → - L1 ⪤*[R, ⓕ{I}V.T] L2. -#R #I #L1 #L2 #V #T * #f1 #HV #Hf1 * #f2 #HT #Hf2 elim (sor_isfin_ex f1 f2) -/3 width=7 by frees_fwd_isfin, frees_flat, lexs_join, ex2_intro/ -qed. - -theorem lfxs_bind_void: ∀R,p,I,L1,L2,V,T. - L1 ⪤*[R, V] L2 → L1.ⓧ ⪤*[R, T] L2.ⓧ → - L1 ⪤*[R, ⓑ{p,I}V.T] L2. -#R #p #I #L1 #L2 #V #T * #f1 #HV #Hf1 * #f2 #HT #Hf2 -lapply (lexs_fwd_bind … Hf2) -Hf2 #Hf2 elim (sor_isfin_ex f1 (⫱f2)) -/3 width=7 by frees_fwd_isfin, frees_bind_void, lexs_join, isfin_tl, ex2_intro/ -qed. - -(* Negated inversion lemmas *************************************************) - -(* Basic_2A1: uses: nllpx_sn_inv_bind nllpx_sn_inv_bind_O *) -lemma lfnxs_inv_bind: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) → - ∀p,I,L1,L2,V,T. (L1 ⪤*[R, ⓑ{p,I}V.T] L2 → ⊥) → - (L1 ⪤*[R, V] L2 → ⊥) ∨ (L1.ⓑ{I}V ⪤*[R, T] L2.ⓑ{I}V → ⊥). -#R #HR #p #I #L1 #L2 #V #T #H elim (lfxs_dec … HR L1 L2 V) -/4 width=2 by lfxs_bind, or_intror, or_introl/ -qed-. - -(* Basic_2A1: uses: nllpx_sn_inv_flat *) -lemma lfnxs_inv_flat: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) → - ∀I,L1,L2,V,T. (L1 ⪤*[R, ⓕ{I}V.T] L2 → ⊥) → - (L1 ⪤*[R, V] L2 → ⊥) ∨ (L1 ⪤*[R, T] L2 → ⊥). -#R #HR #I #L1 #L2 #V #T #H elim (lfxs_dec … HR L1 L2 V) -/4 width=1 by lfxs_flat, or_intror, or_introl/ -qed-. - -lemma lfnxs_inv_bind_void: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) → - ∀p,I,L1,L2,V,T. (L1 ⪤*[R, ⓑ{p,I}V.T] L2 → ⊥) → - (L1 ⪤*[R, V] L2 → ⊥) ∨ (L1.ⓧ ⪤*[R, T] L2.ⓧ → ⊥). -#R #HR #p #I #L1 #L2 #V #T #H elim (lfxs_dec … HR L1 L2 V) -/4 width=2 by lfxs_bind_void, or_intror, or_introl/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/rdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/static/rdeq.ma new file mode 100644 index 000000000..efbd69ae1 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/rdeq.ma @@ -0,0 +1,194 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/notation/relations/stareqsn_5.ma". +include "basic_2/syntax/tdeq_ext.ma". +include "basic_2/static/rex.ma". + +(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******) + +definition rdeq (h) (o): relation3 term lenv lenv ≝ + rex (cdeq h o). + +interpretation + "degree-based equivalence on referred entries (local environment)" + 'StarEqSn h o T L1 L2 = (rdeq h o T L1 L2). + +interpretation + "degree-based ranged equivalence (local environment)" + 'StarEqSn h o f L1 L2 = (sex (cdeq_ext h o) cfull f L1 L2). + +(* Basic properties ***********************************************************) + +lemma frees_tdeq_conf_rdeq (h) (o): ∀f,L1,T1. L1 ⊢ 𝐅*⦃T1⦄ ≘ f → ∀T2. T1 ≛[h, o] T2 → + ∀L2. L1 ≛[h, o, f] L2 → L2 ⊢ 𝐅*⦃T2⦄ ≘ f. +#h #o #f #L1 #T1 #H elim H -f -L1 -T1 +[ #f #L1 #s1 #Hf #X #H1 #L2 #_ + elim (tdeq_inv_sort1 … H1) -H1 #s2 #d #_ #_ #H destruct + /2 width=3 by frees_sort/ +| #f #i #Hf #X #H1 + >(tdeq_inv_lref1 … H1) -X #Y #H2 + >(sex_inv_atom1 … H2) -Y + /2 width=1 by frees_atom/ +| #f #I #L1 #V1 #_ #IH #X #H1 + >(tdeq_inv_lref1 … H1) -X #Y #H2 + elim (sex_inv_next1 … H2) -H2 #Z #L2 #HL12 #HZ #H destruct + elim (ext2_inv_pair_sn … HZ) -HZ #V2 #HV12 #H destruct + /3 width=1 by frees_pair/ +| #f #I #L1 #Hf #X #H1 + >(tdeq_inv_lref1 … H1) -X #Y #H2 + elim (sex_inv_next1 … H2) -H2 #Z #L2 #_ #HZ #H destruct + >(ext2_inv_unit_sn … HZ) -Z /2 width=1 by frees_unit/ +| #f #I #L1 #i #_ #IH #X #H1 + >(tdeq_inv_lref1 … H1) -X #Y #H2 + elim (sex_inv_push1 … H2) -H2 #J #L2 #HL12 #_ #H destruct + /3 width=1 by frees_lref/ +| #f #L1 #l #Hf #X #H1 #L2 #_ + >(tdeq_inv_gref1 … H1) -X /2 width=1 by frees_gref/ +| #f1V #f1T #f1 #p #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #X #H1 + elim (tdeq_inv_pair1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H1 #L2 #HL12 destruct + /6 width=5 by frees_bind, sex_inv_tl, ext2_pair, sle_sex_trans, sor_inv_sle_dx, sor_inv_sle_sn/ +| #f1V #f1T #f1 #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #X #H1 + elim (tdeq_inv_pair1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H1 #L2 #HL12 destruct + /5 width=5 by frees_flat, sle_sex_trans, sor_inv_sle_dx, sor_inv_sle_sn/ +] +qed-. + +lemma frees_tdeq_conf (h) (o): ∀f,L,T1. L ⊢ 𝐅*⦃T1⦄ ≘ f → + ∀T2. T1 ≛[h, o] T2 → L ⊢ 𝐅*⦃T2⦄ ≘ f. +/4 width=7 by frees_tdeq_conf_rdeq, sex_refl, ext2_refl/ qed-. + +lemma frees_rdeq_conf (h) (o): ∀f,L1,T. L1 ⊢ 𝐅*⦃T⦄ ≘ f → + ∀L2. L1 ≛[h, o, f] L2 → L2 ⊢ 𝐅*⦃T⦄ ≘ f. +/2 width=7 by frees_tdeq_conf_rdeq, tdeq_refl/ qed-. + +lemma tdeq_rex_conf (R) (h) (o): s_r_confluent1 … (cdeq h o) (rex R). +#R #h #o #L1 #T1 #T2 #HT12 #L2 * +/3 width=5 by frees_tdeq_conf, ex2_intro/ +qed-. + +lemma tdeq_rex_div (R) (h) (o): ∀T1,T2. T1 ≛[h, o] T2 → + ∀L1,L2. L1 ⪤[R, T2] L2 → L1 ⪤[R, T1] L2. +/3 width=5 by tdeq_rex_conf, tdeq_sym/ qed-. + +lemma tdeq_rdeq_conf (h) (o): s_r_confluent1 … (cdeq h o) (rdeq h o). +/2 width=5 by tdeq_rex_conf/ qed-. + +lemma tdeq_rdeq_div (h) (o): ∀T1,T2. T1 ≛[h, o] T2 → + ∀L1,L2. L1 ≛[h, o, T2] L2 → L1 ≛[h, o, T1] L2. +/2 width=5 by tdeq_rex_div/ qed-. + +lemma rdeq_atom (h) (o): ∀I. ⋆ ≛[h, o, ⓪{I}] ⋆. +/2 width=1 by rex_atom/ qed. + +lemma rdeq_sort (h) (o): ∀I1,I2,L1,L2,s. + L1 ≛[h, o, ⋆s] L2 → L1.ⓘ{I1} ≛[h, o, ⋆s] L2.ⓘ{I2}. +/2 width=1 by rex_sort/ qed. + +lemma rdeq_pair (h) (o): ∀I,L1,L2,V1,V2. L1 ≛[h, o, V1] L2 → V1 ≛[h, o] V2 → + L1.ⓑ{I}V1 ≛[h, o, #0] L2.ⓑ{I}V2. +/2 width=1 by rex_pair/ qed. +(* +lemma rdeq_unit (h) (o): ∀f,I,L1,L2. 𝐈⦃f⦄ → L1 ⪤[cdeq_ext h o, cfull, f] L2 → + L1.ⓤ{I} ≛[h, o, #0] L2.ⓤ{I}. +/2 width=3 by rex_unit/ qed. +*) +lemma rdeq_lref (h) (o): ∀I1,I2,L1,L2,i. + L1 ≛[h, o, #i] L2 → L1.ⓘ{I1} ≛[h, o, #↑i] L2.ⓘ{I2}. +/2 width=1 by rex_lref/ qed. + +lemma rdeq_gref (h) (o): ∀I1,I2,L1,L2,l. + L1 ≛[h, o, §l] L2 → L1.ⓘ{I1} ≛[h, o, §l] L2.ⓘ{I2}. +/2 width=1 by rex_gref/ qed. + +lemma rdeq_bind_repl_dx (h) (o): ∀I,I1,L1,L2.∀T:term. + L1.ⓘ{I} ≛[h, o, T] L2.ⓘ{I1} → + ∀I2. I ≛[h, o] I2 → + L1.ⓘ{I} ≛[h, o, T] L2.ⓘ{I2}. +/2 width=2 by rex_bind_repl_dx/ qed-. + +(* Basic inversion lemmas ***************************************************) + +lemma rdeq_inv_atom_sn (h) (o): ∀Y2. ∀T:term. ⋆ ≛[h, o, T] Y2 → Y2 = ⋆. +/2 width=3 by rex_inv_atom_sn/ qed-. + +lemma rdeq_inv_atom_dx (h) (o): ∀Y1. ∀T:term. Y1 ≛[h, o, T] ⋆ → Y1 = ⋆. +/2 width=3 by rex_inv_atom_dx/ qed-. +(* +lemma rdeq_inv_zero (h) (o): ∀Y1,Y2. Y1 ≛[h, o, #0] Y2 → + ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆ + | ∃∃I,L1,L2,V1,V2. L1 ≛[h, o, V1] L2 & V1 ≛[h, o] V2 & + Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2 + | ∃∃f,I,L1,L2. 𝐈⦃f⦄ & L1 ⪤[cdeq_ext h o, cfull, f] L2 & + Y1 = L1.ⓤ{I} & Y2 = L2.ⓤ{I}. +#h #o #Y1 #Y2 #H elim (rex_inv_zero … H) -H * +/3 width=9 by or3_intro0, or3_intro1, or3_intro2, ex4_5_intro, ex4_4_intro, conj/ +qed-. +*) +lemma rdeq_inv_lref (h) (o): ∀Y1,Y2,i. Y1 ≛[h, o, #↑i] Y2 → + ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆ + | ∃∃I1,I2,L1,L2. L1 ≛[h, o, #i] L2 & + Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. +/2 width=1 by rex_inv_lref/ qed-. + +(* Basic_2A1: uses: lleq_inv_bind lleq_inv_bind_O *) +lemma rdeq_inv_bind (h) (o): ∀p,I,L1,L2,V,T. L1 ≛[h, o, ⓑ{p,I}V.T] L2 → + ∧∧ L1 ≛[h, o, V] L2 & L1.ⓑ{I}V ≛[h, o, T] L2.ⓑ{I}V. +/2 width=2 by rex_inv_bind/ qed-. + +(* Basic_2A1: uses: lleq_inv_flat *) +lemma rdeq_inv_flat (h) (o): ∀I,L1,L2,V,T. L1 ≛[h, o, ⓕ{I}V.T] L2 → + ∧∧ L1 ≛[h, o, V] L2 & L1 ≛[h, o, T] L2. +/2 width=2 by rex_inv_flat/ qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma rdeq_inv_zero_pair_sn (h) (o): ∀I,Y2,L1,V1. L1.ⓑ{I}V1 ≛[h, o, #0] Y2 → + ∃∃L2,V2. L1 ≛[h, o, V1] L2 & V1 ≛[h, o] V2 & Y2 = L2.ⓑ{I}V2. +/2 width=1 by rex_inv_zero_pair_sn/ qed-. + +lemma rdeq_inv_zero_pair_dx (h) (o): ∀I,Y1,L2,V2. Y1 ≛[h, o, #0] L2.ⓑ{I}V2 → + ∃∃L1,V1. L1 ≛[h, o, V1] L2 & V1 ≛[h, o] V2 & Y1 = L1.ⓑ{I}V1. +/2 width=1 by rex_inv_zero_pair_dx/ qed-. + +lemma rdeq_inv_lref_bind_sn (h) (o): ∀I1,Y2,L1,i. L1.ⓘ{I1} ≛[h, o, #↑i] Y2 → + ∃∃I2,L2. L1 ≛[h, o, #i] L2 & Y2 = L2.ⓘ{I2}. +/2 width=2 by rex_inv_lref_bind_sn/ qed-. + +lemma rdeq_inv_lref_bind_dx (h) (o): ∀I2,Y1,L2,i. Y1 ≛[h, o, #↑i] L2.ⓘ{I2} → + ∃∃I1,L1. L1 ≛[h, o, #i] L2 & Y1 = L1.ⓘ{I1}. +/2 width=2 by rex_inv_lref_bind_dx/ qed-. + +(* Basic forward lemmas *****************************************************) + +lemma rdeq_fwd_zero_pair (h) (o): ∀I,K1,K2,V1,V2. + K1.ⓑ{I}V1 ≛[h, o, #0] K2.ⓑ{I}V2 → K1 ≛[h, o, V1] K2. +/2 width=3 by rex_fwd_zero_pair/ qed-. + +(* Basic_2A1: uses: lleq_fwd_bind_sn lleq_fwd_flat_sn *) +lemma rdeq_fwd_pair_sn (h) (o): ∀I,L1,L2,V,T. L1 ≛[h, o, ②{I}V.T] L2 → L1 ≛[h, o, V] L2. +/2 width=3 by rex_fwd_pair_sn/ qed-. + +(* Basic_2A1: uses: lleq_fwd_bind_dx lleq_fwd_bind_O_dx *) +lemma rdeq_fwd_bind_dx (h) (o): ∀p,I,L1,L2,V,T. + L1 ≛[h, o, ⓑ{p,I}V.T] L2 → L1.ⓑ{I}V ≛[h, o, T] L2.ⓑ{I}V. +/2 width=2 by rex_fwd_bind_dx/ qed-. + +(* Basic_2A1: uses: lleq_fwd_flat_dx *) +lemma rdeq_fwd_flat_dx (h) (o): ∀I,L1,L2,V,T. L1 ≛[h, o, ⓕ{I}V.T] L2 → L1 ≛[h, o, T] L2. +/2 width=3 by rex_fwd_flat_dx/ qed-. + +lemma rdeq_fwd_dx (h) (o): ∀I2,L1,K2. ∀T:term. L1 ≛[h, o, T] K2.ⓘ{I2} → + ∃∃I1,K1. L1 = K1.ⓘ{I1}. +/2 width=5 by rex_fwd_dx/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_drops.ma new file mode 100644 index 000000000..1cff3b95c --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_drops.ma @@ -0,0 +1,47 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/lifts_tdeq.ma". +include "basic_2/static/rex_drops.ma". +include "basic_2/static/rdeq.ma". + +(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******) + +(* Properties with generic slicing for local environments *******************) + +lemma rdeq_lifts_sn: ∀h,o. f_dedropable_sn (cdeq h o). +/3 width=5 by rex_liftable_dedropable_sn, tdeq_lifts_sn/ qed-. + +(* Inversion lemmas with generic slicing for local environments *************) + +lemma rdeq_inv_lifts_sn: ∀h,o. f_dropable_sn (cdeq h o). +/2 width=5 by rex_dropable_sn/ qed-. + +(* Note: missing in basic_2A1 *) +lemma rdeq_inv_lifts_dx: ∀h,o. f_dropable_dx (cdeq h o). +/2 width=5 by rex_dropable_dx/ qed-. + +(* Basic_2A1: uses: lleq_inv_lift_le lleq_inv_lift_be lleq_inv_lift_ge *) +lemma rdeq_inv_lifts_bi: ∀h,o,L1,L2,U. L1 ≛[h, o, U] L2 → ∀b,f. 𝐔⦃f⦄ → + ∀K1,K2. ⬇*[b, f] L1 ≘ K1 → ⬇*[b, f] L2 ≘ K2 → + ∀T. ⬆*[f] T ≘ U → K1 ≛[h, o, T] K2. +/2 width=10 by rex_inv_lifts_bi/ qed-. + +lemma rdeq_inv_lref_pair_sn: ∀h,o,L1,L2,i. L1 ≛[h, o, #i] L2 → ∀I,K1,V1. ⬇*[i] L1 ≘ K1.ⓑ{I}V1 → + ∃∃K2,V2. ⬇*[i] L2 ≘ K2.ⓑ{I}V2 & K1 ≛[h, o, V1] K2 & V1 ≛[h, o] V2. +/2 width=3 by rex_inv_lref_pair_sn/ qed-. + +lemma rdeq_inv_lref_pair_dx: ∀h,o,L1,L2,i. L1 ≛[h, o, #i] L2 → ∀I,K2,V2. ⬇*[i] L2 ≘ K2.ⓑ{I}V2 → + ∃∃K1,V1. ⬇*[i] L1 ≘ K1.ⓑ{I}V1 & K1 ≛[h, o, V1] K2 & V1 ≛[h, o] V2. +/2 width=3 by rex_inv_lref_pair_dx/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_fqup.ma b/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_fqup.ma new file mode 100644 index 000000000..5ee09e1c8 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_fqup.ma @@ -0,0 +1,39 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rex_fqup.ma". +include "basic_2/static/rdeq.ma". + +(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******) + +(* Advanced properties ******************************************************) + +lemma rdeq_refl: ∀h,o,T. reflexive … (rdeq h o T). +/2 width=1 by rex_refl/ qed. + +lemma rdeq_pair_refl: ∀h,o,V1,V2. V1 ≛[h, o] V2 → + ∀I,L. ∀T:term. L.ⓑ{I}V1 ≛[h, o, T] L.ⓑ{I}V2. +/2 width=1 by rex_pair_refl/ qed. + +(* Advanced inversion lemmas ************************************************) + +lemma rdeq_inv_bind_void: ∀h,o,p,I,L1,L2,V,T. L1 ≛[h, o, ⓑ{p,I}V.T] L2 → + L1 ≛[h, o, V] L2 ∧ L1.ⓧ ≛[h, o, T] L2.ⓧ. +/2 width=3 by rex_inv_bind_void/ qed-. + +(* Advanced forward lemmas **************************************************) + +lemma rdeq_fwd_bind_dx_void: ∀h,o,p,I,L1,L2,V,T. + L1 ≛[h, o, ⓑ{p,I}V.T] L2 → L1.ⓧ ≛[h, o, T] L2.ⓧ. +/2 width=4 by rex_fwd_bind_dx_void/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_fqus.ma b/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_fqus.ma new file mode 100644 index 000000000..8fe8c8f89 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_fqus.ma @@ -0,0 +1,156 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/s_computation/fqus_fqup.ma". +include "basic_2/static/rdeq_drops.ma". +include "basic_2/static/rdeq_fqup.ma". +include "basic_2/static/rdeq_rdeq.ma". + +(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******) + +(* Properties with extended structural successor for closures ***************) + +lemma fqu_tdeq_conf: ∀h,o,b,G1,G2,L1,L2,U1,T1. ⦃G1, L1, U1⦄ ⊐[b] ⦃G2, L2, T1⦄ → + ∀U2. U1 ≛[h, o] U2 → + ∃∃L,T2. ⦃G1, L1, U2⦄ ⊐[b] ⦃G2, L, T2⦄ & L2 ≛[h, o, T1] L & T1 ≛[h, o] T2. +#h #o #b #G1 #G2 #L1 #L2 #U1 #T1 #H elim H -G1 -G2 -L1 -L2 -U1 -T1 +[ #I #G #L #W #X #H >(tdeq_inv_lref1 … H) -X + /2 width=5 by fqu_lref_O, ex3_2_intro/ +| #I #G #L #W1 #U1 #X #H + elim (tdeq_inv_pair1 … H) -H #W2 #U2 #HW12 #_ #H destruct + /2 width=5 by fqu_pair_sn, ex3_2_intro/ +| #p #I #G #L #W1 #U1 #X #H + elim (tdeq_inv_pair1 … H) -H #W2 #U2 #HW12 #HU12 #H destruct + /3 width=5 by rdeq_pair_refl, fqu_bind_dx, ex3_2_intro/ +| #p #I #G #L #W1 #U1 #Hb #X #H + elim (tdeq_inv_pair1 … H) -H #W2 #U2 #HW12 #HU12 #H destruct + /3 width=5 by fqu_clear, ex3_2_intro/ +| #I #G #L #W1 #U1 #X #H + elim (tdeq_inv_pair1 … H) -H #W2 #U2 #_ #HU12 #H destruct + /2 width=5 by fqu_flat_dx, ex3_2_intro/ +| #I #G #L #T1 #U1 #HTU1 #U2 #HU12 + elim (tdeq_inv_lifts_sn … HU12 … HTU1) -U1 + /3 width=5 by fqu_drop, ex3_2_intro/ +] +qed-. + +lemma tdeq_fqu_trans: ∀h,o,b,G1,G2,L1,L2,U1,T1. ⦃G1, L1, U1⦄ ⊐[b] ⦃G2, L2, T1⦄ → + ∀U2. U2 ≛[h, o] U1 → + ∃∃L,T2. ⦃G1, L1, U2⦄ ⊐[b] ⦃G2, L, T2⦄ & T2 ≛[h, o] T1 & L ≛[h, o, T1] L2. +#h #o #b #G1 #G2 #L1 #L2 #U1 #T1 #H12 #U2 #HU21 +elim (fqu_tdeq_conf … o … H12 U2) -H12 +/3 width=5 by rdeq_sym, tdeq_sym, ex3_2_intro/ +qed-. + +(* Basic_2A1: uses: lleq_fqu_trans *) +lemma rdeq_fqu_trans: ∀h,o,b,G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊐[b] ⦃G2, K2, U⦄ → + ∀L1. L1 ≛[h, o, T] L2 → + ∃∃K1,U0. ⦃G1, L1, T⦄ ⊐[b] ⦃G2, K1, U0⦄ & U0 ≛[h, o] U & K1 ≛[h, o, U] K2. +#h #o #b #G1 #G2 #L2 #K2 #T #U #H elim H -G1 -G2 -L2 -K2 -T -U +[ #I #G #L2 #V2 #L1 #H elim (rdeq_inv_zero_pair_dx … H) -H + #K1 #V1 #HV1 #HV12 #H destruct + /3 width=7 by tdeq_rdeq_conf, fqu_lref_O, ex3_2_intro/ +| * [ #p ] #I #G #L2 #V #T #L1 #H + [ elim (rdeq_inv_bind … H) + | elim (rdeq_inv_flat … H) + ] -H + /2 width=5 by fqu_pair_sn, ex3_2_intro/ +| #p #I #G #L2 #V #T #L1 #H elim (rdeq_inv_bind … H) -H + /2 width=5 by fqu_bind_dx, ex3_2_intro/ +| #p #I #G #L2 #V #T #Hb #L1 #H elim (rdeq_inv_bind_void … H) -H + /3 width=5 by fqu_clear, ex3_2_intro/ +| #I #G #L2 #V #T #L1 #H elim (rdeq_inv_flat … H) -H + /2 width=5 by fqu_flat_dx, ex3_2_intro/ +| #I #G #L2 #T #U #HTU #Y #HU + elim (rdeq_fwd_dx … HU) #L1 #V1 #H destruct + /5 width=14 by rdeq_inv_lifts_bi, fqu_drop, drops_refl, drops_drop, ex3_2_intro/ +] +qed-. + +(* Properties with optional structural successor for closures ***************) + +lemma tdeq_fquq_trans: ∀h,o,b,G1,G2,L1,L2,U1,T1. ⦃G1, L1, U1⦄ ⊐⸮[b] ⦃G2, L2, T1⦄ → + ∀U2. U2 ≛[h, o] U1 → + ∃∃L,T2. ⦃G1, L1, U2⦄ ⊐⸮[b] ⦃G2, L, T2⦄ & T2 ≛[h, o] T1 & L ≛[h, o, T1] L2. +#h #o #b #G1 #G2 #L1 #L2 #U1 #T1 #H elim H -H +[ #H #U2 #HU21 elim (tdeq_fqu_trans … H … HU21) -U1 + /3 width=5 by fqu_fquq, ex3_2_intro/ +| * #HG #HL #HT destruct /2 width=5 by ex3_2_intro/ +] +qed-. + +(* Basic_2A1: was just: lleq_fquq_trans *) +lemma rdeq_fquq_trans: ∀h,o,b,G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊐⸮[b] ⦃G2, K2, U⦄ → + ∀L1. L1 ≛[h, o, T] L2 → + ∃∃K1,U0. ⦃G1, L1, T⦄ ⊐⸮[b] ⦃G2, K1, U0⦄ & U0 ≛[h, o] U & K1 ≛[h, o, U] K2. +#h #o #b #G1 #G2 #L2 #K2 #T #U #H elim H -H +[ #H #L1 #HL12 elim (rdeq_fqu_trans … H … HL12) -L2 /3 width=5 by fqu_fquq, ex3_2_intro/ +| * #HG #HL #HT destruct /2 width=5 by ex3_2_intro/ +] +qed-. + +(* Properties with plus-iterated structural successor for closures **********) + +(* Basic_2A1: was just: lleq_fqup_trans *) +lemma rdeq_fqup_trans: ∀h,o,b,G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊐+[b] ⦃G2, K2, U⦄ → + ∀L1. L1 ≛[h, o, T] L2 → + ∃∃K1,U0. ⦃G1, L1, T⦄ ⊐+[b] ⦃G2, K1, U0⦄ & U0 ≛[h, o] U & K1 ≛[h, o, U] K2. +#h #o #b #G1 #G2 #L2 #K2 #T #U #H @(fqup_ind … H) -G2 -K2 -U +[ #G2 #K2 #U #HTU #L1 #HL12 elim (rdeq_fqu_trans … HTU … HL12) -L2 + /3 width=5 by fqu_fqup, ex3_2_intro/ +| #G #G2 #K #K2 #U #U2 #_ #HU2 #IHTU #L1 #HL12 + elim (IHTU … HL12) -L2 #K0 #U0 #HTU #HU0 #HK0 + elim (rdeq_fqu_trans … HU2 … HK0) -K #K1 #U1 #HU1 #HU12 #HK12 + elim (tdeq_fqu_trans … HU1 … HU0) -U #K3 #U3 #HU03 #HU31 #HK31 + @(ex3_2_intro … K3 U3) (**) (* full auto too slow *) + /3 width=5 by rdeq_trans, tdeq_rdeq_conf, fqup_strap1, tdeq_trans/ +] +qed-. + +lemma tdeq_fqup_trans: ∀h,o,b,G1,G2,L1,L2,U1,T1. ⦃G1, L1, U1⦄ ⊐+[b] ⦃G2, L2, T1⦄ → + ∀U2. U2 ≛[h, o] U1 → + ∃∃L,T2. ⦃G1, L1, U2⦄ ⊐+[b] ⦃G2, L, T2⦄ & T2 ≛[h, o] T1 & L ≛[h, o, T1] L2. +#h #o #b #G1 #G2 #L1 #L2 #U1 #T1 #H @(fqup_ind_dx … H) -G1 -L1 -U1 +[ #G1 #L1 #U1 #H #U2 #HU21 elim (tdeq_fqu_trans … H … HU21) -U1 + /3 width=5 by fqu_fqup, ex3_2_intro/ +| #G1 #G #L1 #L #U1 #U #H #_ #IH #U2 #HU21 + elim (tdeq_fqu_trans … H … HU21) -U1 #L0 #T #H1 #HTU #HL0 + lapply (tdeq_rdeq_div … HTU … HL0) -HL0 #HL0 + elim (IH … HTU) -U #K2 #U1 #H2 #HUT1 #HKL2 + elim (rdeq_fqup_trans … H2 … HL0) -L #K #U #H2 #HU1 #HK2 + lapply (tdeq_rdeq_conf … HUT1 … HK2) -HK2 #HK2 + /3 width=7 by rdeq_trans, fqup_strap2, tdeq_trans, ex3_2_intro/ +] +qed-. + +(* Properties with star-iterated structural successor for closures **********) + +lemma tdeq_fqus_trans: ∀h,o,b,G1,G2,L1,L2,U1,T1. ⦃G1, L1, U1⦄ ⊐*[b] ⦃G2, L2, T1⦄ → + ∀U2. U2 ≛[h, o] U1 → + ∃∃L,T2. ⦃G1, L1, U2⦄ ⊐*[b] ⦃G2, L, T2⦄ & T2 ≛[h, o] T1 & L ≛[h, o, T1] L2. +#h #o #b #G1 #G2 #L1 #L2 #U1 #T1 #H #U2 #HU21 elim(fqus_inv_fqup … H) -H +[ #H elim (tdeq_fqup_trans … H … HU21) -U1 /3 width=5 by fqup_fqus, ex3_2_intro/ +| * #HG #HL #HT destruct /2 width=5 by ex3_2_intro/ +] +qed-. + +(* Basic_2A1: was just: lleq_fqus_trans *) +lemma rdeq_fqus_trans: ∀h,o,b,G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊐*[b] ⦃G2, K2, U⦄ → + ∀L1. L1 ≛[h, o, T] L2 → + ∃∃K1,U0. ⦃G1, L1, T⦄ ⊐*[b] ⦃G2, K1, U0⦄ & U0 ≛[h, o] U & K1 ≛[h, o, U] K2. +#h #o #b #G1 #G2 #L2 #K2 #T #U #H #L1 #HL12 elim(fqus_inv_fqup … H) -H +[ #H elim (rdeq_fqup_trans … H … HL12) -L2 /3 width=5 by fqup_fqus, ex3_2_intro/ +| * #HG #HL #HT destruct /2 width=5 by ex3_2_intro/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_length.ma b/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_length.ma new file mode 100644 index 000000000..a877646c9 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_length.ma @@ -0,0 +1,55 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/lifts_tdeq.ma". +include "basic_2/static/rex_length.ma". +include "basic_2/static/rex_fsle.ma". +include "basic_2/static/rdeq.ma". + +(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******) + +(* Advanved properties with free variables inclusion ************************) + +lemma rdeq_fsge_comp (h) (o): rex_fsge_compatible (cdeq h o). +#h #o #L1 #L2 #T * #f1 #Hf1 #HL12 +lapply (frees_rdeq_conf h o … Hf1 … HL12) +lapply (sex_fwd_length … HL12) +/3 width=8 by lveq_length_eq, ex4_4_intro/ (**) (* full auto fails *) +qed-. + +(* Properties with length for local environments ****************************) + +(* Basic_2A1: uses: lleq_sort *) +lemma rdeq_sort_length (h) (o): ∀L1,L2. |L1| = |L2| → ∀s. L1 ≛[h, o, ⋆s] L2. +/2 width=1 by rex_sort_length/ qed. + +(* Basic_2A1: uses: lleq_gref *) +lemma rdeq_gref_length (h) (o): ∀L1,L2. |L1| = |L2| → ∀l. L1 ≛[h, o, §l] L2. +/2 width=1 by rex_gref_length/ qed. + +lemma rdeq_unit_length (h) (o): ∀L1,L2. |L1| = |L2| → + ∀I. L1.ⓤ{I} ≛[h, o, #0] L2.ⓤ{I}. +/2 width=1 by rex_unit_length/ qed. + +(* Basic_2A1: uses: lleq_lift_le lleq_lift_ge *) +lemma rdeq_lifts_bi (h) (o): ∀L1,L2. |L1| = |L2| → ∀K1,K2,T. K1 ≛[h, o, T] K2 → + ∀b,f. ⬇*[b, f] L1 ≘ K1 → ⬇*[b, f] L2 ≘ K2 → + ∀U. ⬆*[f] T ≘ U → L1 ≛[h, o, U] L2. +/3 width=9 by rex_lifts_bi, tdeq_lifts_sn/ qed-. + +(* Forward lemmas with length for local environments ************************) + +(* Basic_2A1: lleq_fwd_length *) +lemma rdeq_fwd_length (h) (o): ∀L1,L2. ∀T:term. L1 ≛[h, o, T] L2 → |L1| = |L2|. +/2 width=3 by rex_fwd_length/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_rdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_rdeq.ma new file mode 100644 index 000000000..6bd574d2b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_rdeq.ma @@ -0,0 +1,100 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/syntax/ext2_ext2.ma". +include "basic_2/syntax/tdeq_tdeq.ma". +include "basic_2/static/rdeq_length.ma". + +(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******) + +(* Advanced properties ******************************************************) + +(* Basic_2A1: uses: lleq_sym *) +lemma rdeq_sym: ∀h,o,T. symmetric … (rdeq h o T). +/3 width=3 by rdeq_fsge_comp, rex_sym, tdeq_sym/ qed-. + +(* Basic_2A1: uses: lleq_dec *) +lemma rdeq_dec: ∀h,o,L1,L2. ∀T:term. Decidable (L1 ≛[h, o, T] L2). +/3 width=1 by rex_dec, tdeq_dec/ qed-. + +(* Main properties **********************************************************) + +(* Basic_2A1: uses: lleq_bind lleq_bind_O *) +theorem rdeq_bind: ∀h,o,p,I,L1,L2,V1,V2,T. + L1 ≛[h, o, V1] L2 → L1.ⓑ{I}V1 ≛[h, o, T] L2.ⓑ{I}V2 → + L1 ≛[h, o, ⓑ{p,I}V1.T] L2. +/2 width=2 by rex_bind/ qed. + +(* Basic_2A1: uses: lleq_flat *) +theorem rdeq_flat: ∀h,o,I,L1,L2,V,T. L1 ≛[h, o, V] L2 → L1 ≛[h, o, T] L2 → + L1 ≛[h, o, ⓕ{I}V.T] L2. +/2 width=1 by rex_flat/ qed. + +theorem rdeq_bind_void: ∀h,o,p,I,L1,L2,V,T. + L1 ≛[h, o, V] L2 → L1.ⓧ ≛[h, o, T] L2.ⓧ → + L1 ≛[h, o, ⓑ{p,I}V.T] L2. +/2 width=1 by rex_bind_void/ qed. + +(* Basic_2A1: uses: lleq_trans *) +theorem rdeq_trans: ∀h,o,T. Transitive … (rdeq h o T). +#h #o #T #L1 #L * #f1 #Hf1 #HL1 #L2 * #f2 #Hf2 #HL2 +lapply (frees_tdeq_conf_rdeq … Hf1 T … HL1) // #H0 +lapply (frees_mono … Hf2 … H0) -Hf2 -H0 +/5 width=7 by sex_trans, sex_eq_repl_back, tdeq_trans, ext2_trans, ex2_intro/ +qed-. + +(* Basic_2A1: uses: lleq_canc_sn *) +theorem rdeq_canc_sn: ∀h,o,T. left_cancellable … (rdeq h o T). +/3 width=3 by rdeq_trans, rdeq_sym/ qed-. + +(* Basic_2A1: uses: lleq_canc_dx *) +theorem rdeq_canc_dx: ∀h,o,T. right_cancellable … (rdeq h o T). +/3 width=3 by rdeq_trans, rdeq_sym/ qed-. + +theorem rdeq_repl: ∀h,o,L1,L2. ∀T:term. L1 ≛[h, o, T] L2 → + ∀K1. L1 ≛[h, o, T] K1 → ∀K2. L2 ≛[h, o, T] K2 → K1 ≛[h, o, T] K2. +/3 width=3 by rdeq_canc_sn, rdeq_trans/ qed-. + +(* Negated properties *******************************************************) + +(* Note: auto works with /4 width=8/ so rdeq_canc_sn is preferred **********) +(* Basic_2A1: uses: lleq_nlleq_trans *) +lemma rdeq_rdneq_trans: ∀h,o.∀T:term.∀L1,L. L1 ≛[h, o, T] L → + ∀L2. (L ≛[h, o, T] L2 → ⊥) → (L1 ≛[h, o, T] L2 → ⊥). +/3 width=3 by rdeq_canc_sn/ qed-. + +(* Basic_2A1: uses: nlleq_lleq_div *) +lemma rdneq_rdeq_div: ∀h,o.∀T:term.∀L2,L. L2 ≛[h, o, T] L → + ∀L1. (L1 ≛[h, o, T] L → ⊥) → (L1 ≛[h, o, T] L2 → ⊥). +/3 width=3 by rdeq_trans/ qed-. + +theorem rdneq_rdeq_canc_dx: ∀h,o,L1,L. ∀T:term. (L1 ≛[h, o, T] L → ⊥) → + ∀L2. L2 ≛[h, o, T] L → L1 ≛[h, o, T] L2 → ⊥. +/3 width=3 by rdeq_trans/ qed-. + +(* Negated inversion lemmas *************************************************) + +(* Basic_2A1: uses: nlleq_inv_bind nlleq_inv_bind_O *) +lemma rdneq_inv_bind: ∀h,o,p,I,L1,L2,V,T. (L1 ≛[h, o, ⓑ{p,I}V.T] L2 → ⊥) → + (L1 ≛[h, o, V] L2 → ⊥) ∨ (L1.ⓑ{I}V ≛[h, o, T] L2.ⓑ{I}V → ⊥). +/3 width=2 by rnex_inv_bind, tdeq_dec/ qed-. + +(* Basic_2A1: uses: nlleq_inv_flat *) +lemma rdneq_inv_flat: ∀h,o,I,L1,L2,V,T. (L1 ≛[h, o, ⓕ{I}V.T] L2 → ⊥) → + (L1 ≛[h, o, V] L2 → ⊥) ∨ (L1 ≛[h, o, T] L2 → ⊥). +/3 width=2 by rnex_inv_flat, tdeq_dec/ qed-. + +lemma rdneq_inv_bind_void: ∀h,o,p,I,L1,L2,V,T. (L1 ≛[h, o, ⓑ{p,I}V.T] L2 → ⊥) → + (L1 ≛[h, o, V] L2 → ⊥) ∨ (L1.ⓧ ≛[h, o, T] L2.ⓧ → ⊥). +/3 width=3 by rnex_inv_bind_void, tdeq_dec/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_req.ma b/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_req.ma new file mode 100644 index 000000000..46ab07b2b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/rdeq_req.ma @@ -0,0 +1,27 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/req_fsle.ma". +include "basic_2/static/rdeq.ma". + +(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******) + +(* Properties with syntactic equivalence on referred entries ****************) + +lemma req_rdeq: ∀h,o,L1,L2. ∀T:term. L1 ≡[T] L2 → L1 ≛[h, o, T] L2. +/2 width=3 by rex_co/ qed. + +lemma req_rdeq_trans: ∀h,o,L1,L. ∀T:term. L1 ≡[T] L → + ∀L2. L ≛[h, o, T] L2 → L1 ≛[h, o, T] L2. +/2 width=3 by req_rex_trans/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/req.ma b/matita/matita/contribs/lambdadelta/basic_2/static/req.ma new file mode 100644 index 000000000..d01f4ea5d --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/req.ma @@ -0,0 +1,109 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/notation/relations/ideqsn_3.ma". +include "basic_2/static/rex.ma". + +(* SYNTACTIC EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES *********) + +(* Basic_2A1: was: lleq *) +definition req: relation3 term lenv lenv ≝ + rex ceq. + +interpretation + "syntactic equivalence on referred entries (local environment)" + 'IdEqSn T L1 L2 = (req T L1 L2). + +(* Note: "req_transitive R" is equivalent to "rex_transitive ceq R R" *) +(* Basic_2A1: uses: lleq_transitive *) +definition req_transitive: predicate (relation3 lenv term term) ≝ + λR. ∀L2,T1,T2. R L2 T1 T2 → ∀L1. L1 ≡[T1] L2 → R L1 T1 T2. + +(* Basic inversion lemmas ***************************************************) + +lemma req_inv_bind: ∀p,I,L1,L2,V,T. L1 ≡[ⓑ{p,I}V.T] L2 → + ∧∧ L1 ≡[V] L2 & L1.ⓑ{I}V ≡[T] L2.ⓑ{I}V. +/2 width=2 by rex_inv_bind/ qed-. + +lemma req_inv_flat: ∀I,L1,L2,V,T. L1 ≡[ⓕ{I}V.T] L2 → + ∧∧ L1 ≡[V] L2 & L1 ≡[T] L2. +/2 width=2 by rex_inv_flat/ qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma req_inv_zero_pair_sn: ∀I,L2,K1,V. K1.ⓑ{I}V ≡[#0] L2 → + ∃∃K2. K1 ≡[V] K2 & L2 = K2.ⓑ{I}V. +#I #L2 #K1 #V #H +elim (rex_inv_zero_pair_sn … H) -H #K2 #X #HK12 #HX #H destruct +/2 width=3 by ex2_intro/ +qed-. + +lemma req_inv_zero_pair_dx: ∀I,L1,K2,V. L1 ≡[#0] K2.ⓑ{I}V → + ∃∃K1. K1 ≡[V] K2 & L1 = K1.ⓑ{I}V. +#I #L1 #K2 #V #H +elim (rex_inv_zero_pair_dx … H) -H #K1 #X #HK12 #HX #H destruct +/2 width=3 by ex2_intro/ +qed-. + +lemma req_inv_lref_bind_sn: ∀I1,K1,L2,i. K1.ⓘ{I1} ≡[#↑i] L2 → + ∃∃I2,K2. K1 ≡[#i] K2 & L2 = K2.ⓘ{I2}. +/2 width=2 by rex_inv_lref_bind_sn/ qed-. + +lemma req_inv_lref_bind_dx: ∀I2,K2,L1,i. L1 ≡[#↑i] K2.ⓘ{I2} → + ∃∃I1,K1. K1 ≡[#i] K2 & L1 = K1.ⓘ{I1}. +/2 width=2 by rex_inv_lref_bind_dx/ qed-. + +(* Basic forward lemmas *****************************************************) + +(* Basic_2A1: was: llpx_sn_lrefl *) +(* Basic_2A1: this should have been lleq_fwd_llpx_sn *) +lemma req_fwd_rex: ∀R. c_reflexive … R → + ∀L1,L2,T. L1 ≡[T] L2 → L1 ⪤[R, T] L2. +#R #HR #L1 #L2 #T * #f #Hf #HL12 +/4 width=7 by sex_co, cext2_co, ex2_intro/ +qed-. + +(* Basic_properties *********************************************************) + +lemma frees_req_conf: ∀f,L1,T. L1 ⊢ 𝐅*⦃T⦄ ≘ f → + ∀L2. L1 ≡[T] L2 → L2 ⊢ 𝐅*⦃T⦄ ≘ f. +#f #L1 #T #H elim H -f -L1 -T +[ /2 width=3 by frees_sort/ +| #f #i #Hf #L2 #H2 + >(rex_inv_atom_sn … H2) -L2 + /2 width=1 by frees_atom/ +| #f #I #L1 #V1 #_ #IH #Y #H2 + elim (req_inv_zero_pair_sn … H2) -H2 #L2 #HL12 #H destruct + /3 width=1 by frees_pair/ +| #f #I #L1 #Hf #Y #H2 + elim (rex_inv_zero_unit_sn … H2) -H2 #g #L2 #_ #_ #H destruct + /2 width=1 by frees_unit/ +| #f #I #L1 #i #_ #IH #Y #H2 + elim (req_inv_lref_bind_sn … H2) -H2 #J #L2 #HL12 #H destruct + /3 width=1 by frees_lref/ +| /2 width=1 by frees_gref/ +| #f1V #f1T #f1 #p #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #L2 #H2 + elim (req_inv_bind … H2) -H2 /3 width=5 by frees_bind/ +| #f1V #f1T #f1 #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #L2 #H2 + elim (req_inv_flat … H2) -H2 /3 width=5 by frees_flat/ +] +qed-. + +(* Basic_2A1: removed theorems 10: + lleq_ind lleq_fwd_lref + lleq_fwd_drop_sn lleq_fwd_drop_dx + lleq_skip lleq_lref lleq_free + lleq_Y lleq_ge_up lleq_ge + +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/req_fqup.ma b/matita/matita/contribs/lambdadelta/basic_2/static/req_fqup.ma new file mode 100644 index 000000000..3fe7ad753 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/req_fqup.ma @@ -0,0 +1,24 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rex_fqup.ma". +include "basic_2/static/req.ma". + +(* SYNTACTIC EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES *********) + +(* Advanced properties ******************************************************) + +(* Basic_2A1: was: lleq_refl *) +lemma req_refl: ∀T. reflexive … (req T). +/2 width=1 by rex_refl/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/req_fsle.ma b/matita/matita/contribs/lambdadelta/basic_2/static/req_fsle.ma new file mode 100644 index 000000000..7662395c0 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/req_fsle.ma @@ -0,0 +1,33 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/rex_length.ma". +include "basic_2/static/rex_fsle.ma". +include "basic_2/static/req.ma". + +(* SYNTACTIC EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES *********) + +(* Properties with free variables inclusion for restricted closures *********) + +lemma req_fsle_comp: rex_fsle_compatible ceq. +#L1 #L2 #T #HL12 +elim (frees_total L1 T) +/4 width=8 by frees_req_conf, rex_fwd_length, lveq_length_eq, sle_refl, ex4_4_intro/ +qed. + +(* Forward lemmas with free variables inclusion for restricted closures *****) + +lemma req_rex_trans: ∀R. req_transitive R → + ∀L1,L,T. L1 ≡[T] L → ∀L2. L ⪤[R, T] L2 → L1 ⪤[R, T] L2. +/4 width=16 by req_fsle_comp, rex_trans_fsle, rex_trans_next/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/rex.ma b/matita/matita/contribs/lambdadelta/basic_2/static/rex.ma new file mode 100644 index 000000000..cc175f08f --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/rex.ma @@ -0,0 +1,320 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "ground_2/relocation/rtmap_id.ma". +include "basic_2/notation/relations/relation_4.ma". +include "basic_2/syntax/cext2.ma". +include "basic_2/relocation/sex.ma". +include "basic_2/static/frees.ma". + +(* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****) + +definition rex (R) (T): relation lenv ≝ + λL1,L2. ∃∃f. L1 ⊢ 𝐅*⦃T⦄ ≘ f & L1 ⪤[cext2 R, cfull, f] L2. + +interpretation "generic extension on referred entries (local environment)" + 'Relation R T L1 L2 = (rex R T L1 L2). + +definition R_confluent2_rex: relation4 (relation3 lenv term term) + (relation3 lenv term term) … ≝ + λR1,R2,RP1,RP2. + ∀L0,T0,T1. R1 L0 T0 T1 → ∀T2. R2 L0 T0 T2 → + ∀L1. L0 ⪤[RP1, T0] L1 → ∀L2. L0 ⪤[RP2, T0] L2 → + ∃∃T. R2 L1 T1 T & R1 L2 T2 T. + +definition rex_confluent: relation … ≝ + λR1,R2. + ∀K1,K,V1. K1 ⪤[R1, V1] K → ∀V. R1 K1 V1 V → + ∀K2. K ⪤[R2, V] K2 → K ⪤[R2, V1] K2. + +definition rex_transitive: relation3 ? (relation3 ?? term) … ≝ + λR1,R2,R3. + ∀K1,K,V1. K1 ⪤[R1, V1] K → + ∀V. R1 K1 V1 V → ∀V2. R2 K V V2 → R3 K1 V1 V2. + +(* Basic inversion lemmas ***************************************************) + +lemma rex_inv_atom_sn (R): ∀Y2,T. ⋆ ⪤[R, T] Y2 → Y2 = ⋆. +#R #Y2 #T * /2 width=4 by sex_inv_atom1/ +qed-. + +lemma rex_inv_atom_dx (R): ∀Y1,T. Y1 ⪤[R, T] ⋆ → Y1 = ⋆. +#R #I #Y1 * /2 width=4 by sex_inv_atom2/ +qed-. + +lemma rex_inv_sort (R): ∀Y1,Y2,s. Y1 ⪤[R, ⋆s] Y2 → + ∨∨ Y1 = ⋆ ∧ Y2 = ⋆ + | ∃∃I1,I2,L1,L2. L1 ⪤[R, ⋆s] L2 & + Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. +#R * [ | #Y1 #I1 ] #Y2 #s * #f #H1 #H2 +[ lapply (sex_inv_atom1 … H2) -H2 /3 width=1 by or_introl, conj/ +| lapply (frees_inv_sort … H1) -H1 #Hf + elim (isid_inv_gen … Hf) -Hf #g #Hg #H destruct + elim (sex_inv_push1 … H2) -H2 #I2 #L2 #H12 #_ #H destruct + /5 width=7 by frees_sort, ex3_4_intro, ex2_intro, or_intror/ +] +qed-. + +lemma rex_inv_zero (R): ∀Y1,Y2. Y1 ⪤[R, #0] Y2 → + ∨∨ Y1 = ⋆ ∧ Y2 = ⋆ + | ∃∃I,L1,L2,V1,V2. L1 ⪤[R, V1] L2 & R L1 V1 V2 & + Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2 + | ∃∃f,I,L1,L2. 𝐈⦃f⦄ & L1 ⪤[cext2 R, cfull, f] L2 & + Y1 = L1.ⓤ{I} & Y2 = L2.ⓤ{I}. +#R * [ | #Y1 * #I1 [ | #X ] ] #Y2 * #f #H1 #H2 +[ lapply (sex_inv_atom1 … H2) -H2 /3 width=1 by or3_intro0, conj/ +| elim (frees_inv_unit … H1) -H1 #g #HX #H destruct + elim (sex_inv_next1 … H2) -H2 #I2 #L2 #HL12 #H #H2 destruct + >(ext2_inv_unit_sn … H) -H /3 width=8 by or3_intro2, ex4_4_intro/ +| elim (frees_inv_pair … H1) -H1 #g #Hg #H destruct + elim (sex_inv_next1 … H2) -H2 #Z2 #L2 #HL12 #H + elim (ext2_inv_pair_sn … H) -H + /4 width=9 by or3_intro1, ex4_5_intro, ex2_intro/ +] +qed-. + +lemma rex_inv_lref (R): ∀Y1,Y2,i. Y1 ⪤[R, #↑i] Y2 → + ∨∨ Y1 = ⋆ ∧ Y2 = ⋆ + | ∃∃I1,I2,L1,L2. L1 ⪤[R, #i] L2 & + Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. +#R * [ | #Y1 #I1 ] #Y2 #i * #f #H1 #H2 +[ lapply (sex_inv_atom1 … H2) -H2 /3 width=1 by or_introl, conj/ +| elim (frees_inv_lref … H1) -H1 #g #Hg #H destruct + elim (sex_inv_push1 … H2) -H2 + /4 width=7 by ex3_4_intro, ex2_intro, or_intror/ +] +qed-. + +lemma rex_inv_gref (R): ∀Y1,Y2,l. Y1 ⪤[R, §l] Y2 → + ∨∨ Y1 = ⋆ ∧ Y2 = ⋆ + | ∃∃I1,I2,L1,L2. L1 ⪤[R, §l] L2 & + Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}. +#R * [ | #Y1 #I1 ] #Y2 #l * #f #H1 #H2 +[ lapply (sex_inv_atom1 … H2) -H2 /3 width=1 by or_introl, conj/ +| lapply (frees_inv_gref … H1) -H1 #Hf + elim (isid_inv_gen … Hf) -Hf #g #Hg #H destruct + elim (sex_inv_push1 … H2) -H2 #I2 #L2 #H12 #_ #H destruct + /5 width=7 by frees_gref, ex3_4_intro, ex2_intro, or_intror/ +] +qed-. + +(* Basic_2A1: uses: llpx_sn_inv_bind llpx_sn_inv_bind_O *) +lemma rex_inv_bind (R): ∀p,I,L1,L2,V1,V2,T. L1 ⪤[R, ⓑ{p,I}V1.T] L2 → R L1 V1 V2 → + ∧∧ L1 ⪤[R, V1] L2 & L1.ⓑ{I}V1 ⪤[R, T] L2.ⓑ{I}V2. +#R #p #I #L1 #L2 #V1 #V2 #T * #f #Hf #HL #HV elim (frees_inv_bind … Hf) -Hf +/6 width=6 by sle_sex_trans, sex_inv_tl, ext2_pair, sor_inv_sle_dx, sor_inv_sle_sn, ex2_intro, conj/ +qed-. + +(* Basic_2A1: uses: llpx_sn_inv_flat *) +lemma rex_inv_flat (R): ∀I,L1,L2,V,T. L1 ⪤[R, ⓕ{I}V.T] L2 → + ∧∧ L1 ⪤[R, V] L2 & L1 ⪤[R, T] L2. +#R #I #L1 #L2 #V #T * #f #Hf #HL elim (frees_inv_flat … Hf) -Hf +/5 width=6 by sle_sex_trans, sor_inv_sle_dx, sor_inv_sle_sn, ex2_intro, conj/ +qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma rex_inv_sort_bind_sn (R): ∀I1,K1,L2,s. K1.ⓘ{I1} ⪤[R, ⋆s] L2 → + ∃∃I2,K2. K1 ⪤[R, ⋆s] K2 & L2 = K2.ⓘ{I2}. +#R #I1 #K1 #L2 #s #H elim (rex_inv_sort … H) -H * +[ #H destruct +| #Z1 #I2 #Y1 #K2 #Hs #H1 #H2 destruct /2 width=4 by ex2_2_intro/ +] +qed-. + +lemma rex_inv_sort_bind_dx (R): ∀I2,K2,L1,s. L1 ⪤[R, ⋆s] K2.ⓘ{I2} → + ∃∃I1,K1. K1 ⪤[R, ⋆s] K2 & L1 = K1.ⓘ{I1}. +#R #I2 #K2 #L1 #s #H elim (rex_inv_sort … H) -H * +[ #_ #H destruct +| #I1 #Z2 #K1 #Y2 #Hs #H1 #H2 destruct /2 width=4 by ex2_2_intro/ +] +qed-. + +lemma rex_inv_zero_pair_sn (R): ∀I,L2,K1,V1. K1.ⓑ{I}V1 ⪤[R, #0] L2 → + ∃∃K2,V2. K1 ⪤[R, V1] K2 & R K1 V1 V2 & + L2 = K2.ⓑ{I}V2. +#R #I #L2 #K1 #V1 #H elim (rex_inv_zero … H) -H * +[ #H destruct +| #Z #Y1 #K2 #X1 #V2 #HK12 #HV12 #H1 #H2 destruct + /2 width=5 by ex3_2_intro/ +| #f #Z #Y1 #Y2 #_ #_ #H destruct +] +qed-. + +lemma rex_inv_zero_pair_dx (R): ∀I,L1,K2,V2. L1 ⪤[R, #0] K2.ⓑ{I}V2 → + ∃∃K1,V1. K1 ⪤[R, V1] K2 & R K1 V1 V2 & + L1 = K1.ⓑ{I}V1. +#R #I #L1 #K2 #V2 #H elim (rex_inv_zero … H) -H * +[ #_ #H destruct +| #Z #K1 #Y2 #V1 #X2 #HK12 #HV12 #H1 #H2 destruct + /2 width=5 by ex3_2_intro/ +| #f #Z #Y1 #Y2 #_ #_ #_ #H destruct +] +qed-. + +lemma rex_inv_zero_unit_sn (R): ∀I,K1,L2. K1.ⓤ{I} ⪤[R, #0] L2 → + ∃∃f,K2. 𝐈⦃f⦄ & K1 ⪤[cext2 R, cfull, f] K2 & + L2 = K2.ⓤ{I}. +#R #I #K1 #L2 #H elim (rex_inv_zero … H) -H * +[ #H destruct +| #Z #Y1 #Y2 #X1 #X2 #_ #_ #H destruct +| #f #Z #Y1 #K2 #Hf #HK12 #H1 #H2 destruct /2 width=5 by ex3_2_intro/ +] +qed-. + +lemma rex_inv_zero_unit_dx (R): ∀I,L1,K2. L1 ⪤[R, #0] K2.ⓤ{I} → + ∃∃f,K1. 𝐈⦃f⦄ & K1 ⪤[cext2 R, cfull, f] K2 & + L1 = K1.ⓤ{I}. +#R #I #L1 #K2 #H elim (rex_inv_zero … H) -H * +[ #_ #H destruct +| #Z #Y1 #Y2 #X1 #X2 #_ #_ #_ #H destruct +| #f #Z #K1 #Y2 #Hf #HK12 #H1 #H2 destruct /2 width=5 by ex3_2_intro/ +] +qed-. + +lemma rex_inv_lref_bind_sn (R): ∀I1,K1,L2,i. K1.ⓘ{I1} ⪤[R, #↑i] L2 → + ∃∃I2,K2. K1 ⪤[R, #i] K2 & L2 = K2.ⓘ{I2}. +#R #I1 #K1 #L2 #i #H elim (rex_inv_lref … H) -H * +[ #H destruct +| #Z1 #I2 #Y1 #K2 #Hi #H1 #H2 destruct /2 width=4 by ex2_2_intro/ +] +qed-. + +lemma rex_inv_lref_bind_dx (R): ∀I2,K2,L1,i. L1 ⪤[R, #↑i] K2.ⓘ{I2} → + ∃∃I1,K1. K1 ⪤[R, #i] K2 & L1 = K1.ⓘ{I1}. +#R #I2 #K2 #L1 #i #H elim (rex_inv_lref … H) -H * +[ #_ #H destruct +| #I1 #Z2 #K1 #Y2 #Hi #H1 #H2 destruct /2 width=4 by ex2_2_intro/ +] +qed-. + +lemma rex_inv_gref_bind_sn (R): ∀I1,K1,L2,l. K1.ⓘ{I1} ⪤[R, §l] L2 → + ∃∃I2,K2. K1 ⪤[R, §l] K2 & L2 = K2.ⓘ{I2}. +#R #I1 #K1 #L2 #l #H elim (rex_inv_gref … H) -H * +[ #H destruct +| #Z1 #I2 #Y1 #K2 #Hl #H1 #H2 destruct /2 width=4 by ex2_2_intro/ +] +qed-. + +lemma rex_inv_gref_bind_dx (R): ∀I2,K2,L1,l. L1 ⪤[R, §l] K2.ⓘ{I2} → + ∃∃I1,K1. K1 ⪤[R, §l] K2 & L1 = K1.ⓘ{I1}. +#R #I2 #K2 #L1 #l #H elim (rex_inv_gref … H) -H * +[ #_ #H destruct +| #I1 #Z2 #K1 #Y2 #Hl #H1 #H2 destruct /2 width=4 by ex2_2_intro/ +] +qed-. + +(* Basic forward lemmas *****************************************************) + +lemma rex_fwd_zero_pair (R): ∀I,K1,K2,V1,V2. + K1.ⓑ{I}V1 ⪤[R, #0] K2.ⓑ{I}V2 → K1 ⪤[R, V1] K2. +#R #I #K1 #K2 #V1 #V2 #H +elim (rex_inv_zero_pair_sn … H) -H #Y #X #HK12 #_ #H destruct // +qed-. + +(* Basic_2A1: uses: llpx_sn_fwd_pair_sn llpx_sn_fwd_bind_sn llpx_sn_fwd_flat_sn *) +lemma rex_fwd_pair_sn (R): ∀I,L1,L2,V,T. L1 ⪤[R, ②{I}V.T] L2 → L1 ⪤[R, V] L2. +#R * [ #p ] #I #L1 #L2 #V #T * #f #Hf #HL +[ elim (frees_inv_bind … Hf) | elim (frees_inv_flat … Hf) ] -Hf +/4 width=6 by sle_sex_trans, sor_inv_sle_sn, ex2_intro/ +qed-. + +(* Basic_2A1: uses: llpx_sn_fwd_bind_dx llpx_sn_fwd_bind_O_dx *) +lemma rex_fwd_bind_dx (R): ∀p,I,L1,L2,V1,V2,T. L1 ⪤[R, ⓑ{p,I}V1.T] L2 → + R L1 V1 V2 → L1.ⓑ{I}V1 ⪤[R, T] L2.ⓑ{I}V2. +#R #p #I #L1 #L2 #V1 #V2 #T #H #HV elim (rex_inv_bind … H HV) -H -HV // +qed-. + +(* Basic_2A1: uses: llpx_sn_fwd_flat_dx *) +lemma rex_fwd_flat_dx (R): ∀I,L1,L2,V,T. L1 ⪤[R, ⓕ{I}V.T] L2 → L1 ⪤[R, T] L2. +#R #I #L1 #L2 #V #T #H elim (rex_inv_flat … H) -H // +qed-. + +lemma rex_fwd_dx (R): ∀I2,L1,K2,T. L1 ⪤[R, T] K2.ⓘ{I2} → + ∃∃I1,K1. L1 = K1.ⓘ{I1}. +#R #I2 #L1 #K2 #T * #f elim (pn_split f) * #g #Hg #_ #Hf destruct +[ elim (sex_inv_push2 … Hf) | elim (sex_inv_next2 … Hf) ] -Hf #I1 #K1 #_ #_ #H destruct +/2 width=3 by ex1_2_intro/ +qed-. + +(* Basic properties *********************************************************) + +lemma rex_atom (R): ∀I. ⋆ ⪤[R, ⓪{I}] ⋆. +#R * /3 width=3 by frees_sort, frees_atom, frees_gref, sex_atom, ex2_intro/ +qed. + +lemma rex_sort (R): ∀I1,I2,L1,L2,s. + L1 ⪤[R, ⋆s] L2 → L1.ⓘ{I1} ⪤[R, ⋆s] L2.ⓘ{I2}. +#R #I1 #I2 #L1 #L2 #s * #f #Hf #H12 +lapply (frees_inv_sort … Hf) -Hf +/4 width=3 by frees_sort, sex_push, isid_push, ex2_intro/ +qed. + +lemma rex_pair (R): ∀I,L1,L2,V1,V2. L1 ⪤[R, V1] L2 → + R L1 V1 V2 → L1.ⓑ{I}V1 ⪤[R, #0] L2.ⓑ{I}V2. +#R #I1 #I2 #L1 #L2 #V1 * +/4 width=3 by ext2_pair, frees_pair, sex_next, ex2_intro/ +qed. + +lemma rex_unit (R): ∀f,I,L1,L2. 𝐈⦃f⦄ → L1 ⪤[cext2 R, cfull, f] L2 → + L1.ⓤ{I} ⪤[R, #0] L2.ⓤ{I}. +/4 width=3 by frees_unit, sex_next, ext2_unit, ex2_intro/ qed. + +lemma rex_lref (R): ∀I1,I2,L1,L2,i. + L1 ⪤[R, #i] L2 → L1.ⓘ{I1} ⪤[R, #↑i] L2.ⓘ{I2}. +#R #I1 #I2 #L1 #L2 #i * /3 width=3 by sex_push, frees_lref, ex2_intro/ +qed. + +lemma rex_gref (R): ∀I1,I2,L1,L2,l. + L1 ⪤[R, §l] L2 → L1.ⓘ{I1} ⪤[R, §l] L2.ⓘ{I2}. +#R #I1 #I2 #L1 #L2 #l * #f #Hf #H12 +lapply (frees_inv_gref … Hf) -Hf +/4 width=3 by frees_gref, sex_push, isid_push, ex2_intro/ +qed. + +lemma rex_bind_repl_dx (R): ∀I,I1,L1,L2,T. + L1.ⓘ{I} ⪤[R, T] L2.ⓘ{I1} → + ∀I2. cext2 R L1 I I2 → + L1.ⓘ{I} ⪤[R, T] L2.ⓘ{I2}. +#R #I #I1 #L1 #L2 #T * #f #Hf #HL12 #I2 #HR +/3 width=5 by sex_pair_repl, ex2_intro/ +qed-. + +(* Basic_2A1: uses: llpx_sn_co *) +lemma rex_co (R1) (R2): (∀L,T1,T2. R1 L T1 T2 → R2 L T1 T2) → + ∀L1,L2,T. L1 ⪤[R1, T] L2 → L1 ⪤[R2, T] L2. +#R1 #R2 #HR #L1 #L2 #T * /5 width=7 by sex_co, cext2_co, ex2_intro/ +qed-. + +lemma rex_isid (R1) (R2): ∀L1,L2,T1,T2. + (∀f. L1 ⊢ 𝐅*⦃T1⦄ ≘ f → 𝐈⦃f⦄) → + (∀f. 𝐈⦃f⦄ → L1 ⊢ 𝐅*⦃T2⦄ ≘ f) → + L1 ⪤[R1, T1] L2 → L1 ⪤[R2, T2] L2. +#R1 #R2 #L1 #L2 #T1 #T2 #H1 #H2 * +/4 width=7 by sex_co_isid, ex2_intro/ +qed-. + +lemma rex_unit_sn (R1) (R2): + ∀I,K1,L2. K1.ⓤ{I} ⪤[R1, #0] L2 → K1.ⓤ{I} ⪤[R2, #0] L2. +#R1 #R2 #I #K1 #L2 #H +elim (rex_inv_zero_unit_sn … H) -H #f #K2 #Hf #HK12 #H destruct +/3 width=7 by rex_unit, sex_co_isid/ +qed-. + +(* Basic_2A1: removed theorems 9: + llpx_sn_skip llpx_sn_lref llpx_sn_free + llpx_sn_fwd_lref + llpx_sn_Y llpx_sn_ge_up llpx_sn_ge + llpx_sn_fwd_drop_sn llpx_sn_fwd_drop_dx +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/rex_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/static/rex_drops.ma new file mode 100644 index 000000000..7ce890273 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/rex_drops.ma @@ -0,0 +1,124 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/drops_cext2.ma". +include "basic_2/relocation/drops_sex.ma". +include "basic_2/static/frees_drops.ma". +include "basic_2/static/rex.ma". + +(* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****) + +definition f_dedropable_sn: predicate (relation3 lenv term term) ≝ + λR. ∀b,f,L1,K1. ⬇*[b, f] L1 ≘ K1 → + ∀K2,T. K1 ⪤[R, T] K2 → ∀U. ⬆*[f] T ≘ U → + ∃∃L2. L1 ⪤[R, U] L2 & ⬇*[b, f] L2 ≘ K2 & L1 ≡[f] L2. + +definition f_dropable_sn: predicate (relation3 lenv term term) ≝ + λR. ∀b,f,L1,K1. ⬇*[b, f] L1 ≘ K1 → 𝐔⦃f⦄ → + ∀L2,U. L1 ⪤[R, U] L2 → ∀T. ⬆*[f] T ≘ U → + ∃∃K2. K1 ⪤[R, T] K2 & ⬇*[b, f] L2 ≘ K2. + +definition f_dropable_dx: predicate (relation3 lenv term term) ≝ + λR. ∀L1,L2,U. L1 ⪤[R, U] L2 → + ∀b,f,K2. ⬇*[b, f] L2 ≘ K2 → 𝐔⦃f⦄ → ∀T. ⬆*[f] T ≘ U → + ∃∃K1. ⬇*[b, f] L1 ≘ K1 & K1 ⪤[R, T] K2. + +definition f_transitive_next: relation3 … ≝ λR1,R2,R3. + ∀f,L,T. L ⊢ 𝐅*⦃T⦄ ≘ f → + ∀g,I,K,n. ⬇*[n] L ≘ K.ⓘ{I} → ↑g = ⫱*[n] f → + sex_transitive (cext2 R1) (cext2 R2) (cext2 R3) (cext2 R1) cfull g K I. + +(* Properties with generic slicing for local environments *******************) + +lemma rex_liftable_dedropable_sn: ∀R. (∀L. reflexive ? (R L)) → + d_liftable2_sn … lifts R → f_dedropable_sn R. +#R #H1R #H2R #b #f #L1 #K1 #HLK1 #K2 #T * #f1 #Hf1 #HK12 #U #HTU +elim (frees_total L1 U) #f2 #Hf2 +lapply (frees_fwd_coafter … Hf2 … HLK1 … HTU … Hf1) -HTU #Hf +elim (sex_liftable_co_dedropable_sn … HLK1 … HK12 … Hf) -f1 -K1 +/3 width=6 by cext2_d_liftable2_sn, cfull_lift_sn, ext2_refl, ex3_intro, ex2_intro/ +qed-. + +lemma rex_trans_next: ∀R1,R2,R3. rex_transitive R1 R2 R3 → f_transitive_next R1 R2 R3. +#R1 #R2 #R3 #HR #f #L1 #T #Hf #g #I1 #K1 #n #HLK #Hgf #I #H +generalize in match HLK; -HLK elim H -I1 -I +[ #I #_ #L2 #_ #I2 #H + lapply (ext2_inv_unit_sn … H) -H #H destruct + /2 width=1 by ext2_unit/ +| #I #V1 #V #HV1 #HLK1 #L2 #HL12 #I2 #H + elim (ext2_inv_pair_sn … H) -H #V2 #HV2 #H destruct + elim (frees_inv_drops_next … Hf … HLK1 … Hgf) -f -HLK1 #f #Hf #Hfg + /5 width=5 by ext2_pair, sle_sex_trans, ex2_intro/ +] +qed. + +(* Inversion lemmas with generic slicing for local environments *************) + +(* Basic_2A1: uses: llpx_sn_inv_lift_le llpx_sn_inv_lift_be llpx_sn_inv_lift_ge *) +(* Basic_2A1: was: llpx_sn_drop_conf_O *) +lemma rex_dropable_sn: ∀R. f_dropable_sn R. +#R #b #f #L1 #K1 #HLK1 #H1f #L2 #U * #f2 #Hf2 #HL12 #T #HTU +elim (frees_total K1 T) #f1 #Hf1 +lapply (frees_fwd_coafter … Hf2 … HLK1 … HTU … Hf1) -HTU #H2f +elim (sex_co_dropable_sn … HLK1 … HL12 … H2f) -f2 -L1 +/3 width=3 by ex2_intro/ +qed-. + +(* Basic_2A1: was: llpx_sn_drop_trans_O *) +(* Note: the proof might be simplified *) +lemma rex_dropable_dx: ∀R. f_dropable_dx R. +#R #L1 #L2 #U * #f2 #Hf2 #HL12 #b #f #K2 #HLK2 #H1f #T #HTU +elim (drops_isuni_ex … H1f L1) #K1 #HLK1 +elim (frees_total K1 T) #f1 #Hf1 +lapply (frees_fwd_coafter … Hf2 … HLK1 … HTU … Hf1) -K1 #H2f +elim (sex_co_dropable_dx … HL12 … HLK2 … H2f) -L2 +/4 width=9 by frees_inv_lifts, ex2_intro/ +qed-. + +(* Basic_2A1: uses: llpx_sn_inv_lift_O *) +lemma rex_inv_lifts_bi: ∀R,L1,L2,U. L1 ⪤[R, U] L2 → ∀b,f. 𝐔⦃f⦄ → + ∀K1,K2. ⬇*[b, f] L1 ≘ K1 → ⬇*[b, f] L2 ≘ K2 → + ∀T. ⬆*[f] T ≘ U → K1 ⪤[R, T] K2. +#R #L1 #L2 #U #HL12 #b #f #Hf #K1 #K2 #HLK1 #HLK2 #T #HTU +elim (rex_dropable_sn … HLK1 … HL12 … HTU) -L1 -U // #Y #HK12 #HY +lapply (drops_mono … HY … HLK2) -b -f -L2 #H destruct // +qed-. + +lemma rex_inv_lref_pair_sn: ∀R,L1,L2,i. L1 ⪤[R, #i] L2 → ∀I,K1,V1. ⬇*[i] L1 ≘ K1.ⓑ{I}V1 → + ∃∃K2,V2. ⬇*[i] L2 ≘ K2.ⓑ{I}V2 & K1 ⪤[R, V1] K2 & R K1 V1 V2. +#R #L1 #L2 #i #HL12 #I #K1 #V1 #HLK1 elim (rex_dropable_sn … HLK1 … HL12 (#0)) -HLK1 -HL12 // +#Y #HY #HLK2 elim (rex_inv_zero_pair_sn … HY) -HY +#K2 #V2 #HK12 #HV12 #H destruct /2 width=5 by ex3_2_intro/ +qed-. + +lemma rex_inv_lref_pair_dx: ∀R,L1,L2,i. L1 ⪤[R, #i] L2 → ∀I,K2,V2. ⬇*[i] L2 ≘ K2.ⓑ{I}V2 → + ∃∃K1,V1. ⬇*[i] L1 ≘ K1.ⓑ{I}V1 & K1 ⪤[R, V1] K2 & R K1 V1 V2. +#R #L1 #L2 #i #HL12 #I #K2 #V2 #HLK2 elim (rex_dropable_dx … HL12 … HLK2 … (#0)) -HLK2 -HL12 // +#Y #HLK1 #HY elim (rex_inv_zero_pair_dx … HY) -HY +#K1 #V1 #HK12 #HV12 #H destruct /2 width=5 by ex3_2_intro/ +qed-. + +lemma rex_inv_lref_unit_sn: ∀R,L1,L2,i. L1 ⪤[R, #i] L2 → ∀I,K1. ⬇*[i] L1 ≘ K1.ⓤ{I} → + ∃∃f,K2. ⬇*[i] L2 ≘ K2.ⓤ{I} & K1 ⪤[cext2 R, cfull, f] K2 & 𝐈⦃f⦄. +#R #L1 #L2 #i #HL12 #I #K1 #HLK1 elim (rex_dropable_sn … HLK1 … HL12 (#0)) -HLK1 -HL12 // +#Y #HY #HLK2 elim (rex_inv_zero_unit_sn … HY) -HY +#f #K2 #Hf #HK12 #H destruct /2 width=5 by ex3_2_intro/ +qed-. + +lemma rex_inv_lref_unit_dx: ∀R,L1,L2,i. L1 ⪤[R, #i] L2 → ∀I,K2. ⬇*[i] L2 ≘ K2.ⓤ{I} → + ∃∃f,K1. ⬇*[i] L1 ≘ K1.ⓤ{I} & K1 ⪤[cext2 R, cfull, f] K2 & 𝐈⦃f⦄. +#R #L1 #L2 #i #HL12 #I #K2 #HLK2 elim (rex_dropable_dx … HL12 … HLK2 … (#0)) -HLK2 -HL12 // +#Y #HLK1 #HY elim (rex_inv_zero_unit_dx … HY) -HY +#f #K2 #Hf #HK12 #H destruct /2 width=5 by ex3_2_intro/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/rex_fqup.ma b/matita/matita/contribs/lambdadelta/basic_2/static/rex_fqup.ma new file mode 100644 index 000000000..6060d3ff6 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/rex_fqup.ma @@ -0,0 +1,49 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/frees_fqup.ma". +include "basic_2/static/rex.ma". + +(* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****) + +(* Advanced properties ******************************************************) + +(* Basic_2A1: uses: llpx_sn_refl *) +lemma rex_refl: ∀R. (∀L. reflexive … (R L)) → ∀L,T. L ⪤[R, T] L. +#R #HR #L #T elim (frees_total L T) +/4 width=3 by sex_refl, ext2_refl, ex2_intro/ +qed. + +lemma rex_pair_refl: ∀R. (∀L. reflexive … (R L)) → + ∀L,V1,V2. R L V1 V2 → ∀I,T. L.ⓑ{I}V1 ⪤[R, T] L.ⓑ{I}V2. +#R #HR #L #V1 #V2 #HV12 #I #T +elim (frees_total (L.ⓑ{I}V1) T) #f #Hf +elim (pn_split f) * #g #H destruct +/5 width=3 by sex_refl, sex_next, sex_push, ext2_refl, ext2_pair, ex2_intro/ +qed. + +(* Advanced inversion lemmas ************************************************) + +lemma rex_inv_bind_void: ∀R,p,I,L1,L2,V,T. L1 ⪤[R, ⓑ{p,I}V.T] L2 → + L1 ⪤[R, V] L2 ∧ L1.ⓧ ⪤[R, T] L2.ⓧ. +#R #p #I #L1 #L2 #V #T * #f #Hf #HL elim (frees_inv_bind_void … Hf) -Hf +/6 width=6 by sle_sex_trans, sex_inv_tl, sor_inv_sle_dx, sor_inv_sle_sn, ex2_intro, conj/ +qed-. + +(* Advanced forward lemmas **************************************************) + +lemma rex_fwd_bind_dx_void: ∀R,p,I,L1,L2,V,T. L1 ⪤[R, ⓑ{p,I}V.T] L2 → + L1.ⓧ ⪤[R, T] L2.ⓧ. +#R #p #I #L1 #L2 #V #T #H elim (rex_inv_bind_void … H) -H // +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/rex_fsle.ma b/matita/matita/contribs/lambdadelta/basic_2/static/rex_fsle.ma new file mode 100644 index 000000000..b69cd2efb --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/rex_fsle.ma @@ -0,0 +1,176 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/sex_length.ma". +include "basic_2/static/fsle_fsle.ma". +include "basic_2/static/rex_drops.ma". +include "basic_2/static/rex_rex.ma". + +(* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****) + +definition R_fsge_compatible: predicate (relation3 …) ≝ λRN. + ∀L,T1,T2. RN L T1 T2 → ⦃L, T2⦄ ⊆ ⦃L, T1⦄. + +definition rex_fsge_compatible: predicate (relation3 …) ≝ λRN. + ∀L1,L2,T. L1 ⪤[RN, T] L2 → ⦃L2, T⦄ ⊆ ⦃L1, T⦄. + +definition rex_fsle_compatible: predicate (relation3 …) ≝ λRN. + ∀L1,L2,T. L1 ⪤[RN, T] L2 → ⦃L1, T⦄ ⊆ ⦃L2, T⦄. + +(* Basic inversions with free variables inclusion for restricted closures ***) + +lemma frees_sex_conf: ∀R. rex_fsge_compatible R → + ∀L1,T,f1. L1 ⊢ 𝐅*⦃T⦄ ≘ f1 → + ∀L2. L1 ⪤[cext2 R, cfull, f1] L2 → + ∃∃f2. L2 ⊢ 𝐅*⦃T⦄ ≘ f2 & f2 ⊆ f1. +#R #HR #L1 #T #f1 #Hf1 #L2 #H1L +lapply (HR L1 L2 T ?) /2 width=3 by ex2_intro/ #H2L +@(fsle_frees_trans_eq … H2L … Hf1) /3 width=4 by sex_fwd_length, sym_eq/ +qed-. + +(* Properties with free variables inclusion for restricted closures *********) + +(* Note: we just need lveq_inv_refl: ∀L,n1,n2. L ≋ⓧ*[n1, n2] L → ∧∧ 0 = n1 & 0 = n2 *) +lemma fsge_rex_trans: ∀R,L1,T1,T2. ⦃L1, T1⦄ ⊆ ⦃L1, T2⦄ → + ∀L2. L1 ⪤[R, T2] L2 → L1 ⪤[R, T1] L2. +#R #L1 #T1 #T2 * #n1 #n2 #f1 #f2 #Hf1 #Hf2 #Hn #Hf #L2 #HL12 +elim (lveq_inj_length … Hn ?) // #H1 #H2 destruct +/4 width=5 by rex_inv_frees, sle_sex_trans, ex2_intro/ +qed-. + +lemma rex_sym: ∀R. rex_fsge_compatible R → + (∀L1,L2,T1,T2. R L1 T1 T2 → R L2 T2 T1) → + ∀T. symmetric … (rex R T). +#R #H1R #H2R #T #L1 #L2 +* #f1 #Hf1 #HL12 +elim (frees_sex_conf … Hf1 … HL12) -Hf1 // +/5 width=5 by sle_sex_trans, sex_sym, cext2_sym, ex2_intro/ +qed-. + +lemma rex_pair_sn_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) → + rex_fsge_compatible R1 → + ∀L1,L2,V. L1 ⪤[R1, V] L2 → ∀I,T. + ∃∃L. L1 ⪤[R1, ②{I}V.T] L & L ⪤[R2, V] L2. +#R1 #R2 #HR1 #HR2 #HR #L1 #L2 #V * #f #Hf #HL12 * [ #p ] #I #T +[ elim (frees_total L1 (ⓑ{p,I}V.T)) #g #Hg + elim (frees_inv_bind … Hg) #y1 #y2 #H #_ #Hy +| elim (frees_total L1 (ⓕ{I}V.T)) #g #Hg + elim (frees_inv_flat … Hg) #y1 #y2 #H #_ #Hy +] +lapply(frees_mono … H … Hf) -H #H1 +lapply (sor_eq_repl_back1 … Hy … H1) -y1 #Hy +lapply (sor_inv_sle_sn … Hy) -y2 #Hfg +elim (sex_sle_split (cext2 R1) (cext2 R2) … HL12 … Hfg) -HL12 /2 width=1 by ext2_refl/ #L #HL1 #HL2 +lapply (sle_sex_trans … HL1 … Hfg) // #H +elim (frees_sex_conf … Hf … H) -Hf -H +/4 width=7 by sle_sex_trans, ex2_intro/ +qed-. + +lemma rex_flat_dx_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) → + rex_fsge_compatible R1 → + ∀L1,L2,T. L1 ⪤[R1, T] L2 → ∀I,V. + ∃∃L. L1 ⪤[R1, ⓕ{I}V.T] L & L ⪤[R2, T] L2. +#R1 #R2 #HR1 #HR2 #HR #L1 #L2 #T * #f #Hf #HL12 #I #V +elim (frees_total L1 (ⓕ{I}V.T)) #g #Hg +elim (frees_inv_flat … Hg) #y1 #y2 #_ #H #Hy +lapply(frees_mono … H … Hf) -H #H2 +lapply (sor_eq_repl_back2 … Hy … H2) -y2 #Hy +lapply (sor_inv_sle_dx … Hy) -y1 #Hfg +elim (sex_sle_split (cext2 R1) (cext2 R2) … HL12 … Hfg) -HL12 /2 width=1 by ext2_refl/ #L #HL1 #HL2 +lapply (sle_sex_trans … HL1 … Hfg) // #H +elim (frees_sex_conf … Hf … H) -Hf -H +/4 width=7 by sle_sex_trans, ex2_intro/ +qed-. + +lemma rex_bind_dx_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) → + rex_fsge_compatible R1 → + ∀I,L1,L2,V1,T. L1.ⓑ{I}V1 ⪤[R1, T] L2 → ∀p. + ∃∃L,V. L1 ⪤[R1, ⓑ{p,I}V1.T] L & L.ⓑ{I}V ⪤[R2, T] L2 & R1 L1 V1 V. +#R1 #R2 #HR1 #HR2 #HR #I #L1 #L2 #V1 #T * #f #Hf #HL12 #p +elim (frees_total L1 (ⓑ{p,I}V1.T)) #g #Hg +elim (frees_inv_bind … Hg) #y1 #y2 #_ #H #Hy +lapply(frees_mono … H … Hf) -H #H2 +lapply (tl_eq_repl … H2) -H2 #H2 +lapply (sor_eq_repl_back2 … Hy … H2) -y2 #Hy +lapply (sor_inv_sle_dx … Hy) -y1 #Hfg +lapply (sle_inv_tl_sn … Hfg) -Hfg #Hfg +elim (sex_sle_split (cext2 R1) (cext2 R2) … HL12 … Hfg) -HL12 /2 width=1 by ext2_refl/ #Y #H #HL2 +lapply (sle_sex_trans … H … Hfg) // #H0 +elim (sex_inv_next1 … H) -H #Z #L #HL1 #H +elim (ext2_inv_pair_sn … H) -H #V #HV #H1 #H2 destruct +elim (frees_sex_conf … Hf … H0) -Hf -H0 +/4 width=7 by sle_sex_trans, ex3_2_intro, ex2_intro/ +qed-. + +lemma rex_bind_dx_split_void: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) → + rex_fsge_compatible R1 → + ∀L1,L2,T. L1.ⓧ ⪤[R1, T] L2 → ∀p,I,V. + ∃∃L. L1 ⪤[R1, ⓑ{p,I}V.T] L & L.ⓧ ⪤[R2, T] L2. +#R1 #R2 #HR1 #HR2 #HR #L1 #L2 #T * #f #Hf #HL12 #p #I #V +elim (frees_total L1 (ⓑ{p,I}V.T)) #g #Hg +elim (frees_inv_bind_void … Hg) #y1 #y2 #_ #H #Hy +lapply(frees_mono … H … Hf) -H #H2 +lapply (tl_eq_repl … H2) -H2 #H2 +lapply (sor_eq_repl_back2 … Hy … H2) -y2 #Hy +lapply (sor_inv_sle_dx … Hy) -y1 #Hfg +lapply (sle_inv_tl_sn … Hfg) -Hfg #Hfg +elim (sex_sle_split (cext2 R1) (cext2 R2) … HL12 … Hfg) -HL12 /2 width=1 by ext2_refl/ #Y #H #HL2 +lapply (sle_sex_trans … H … Hfg) // #H0 +elim (sex_inv_next1 … H) -H #Z #L #HL1 #H +elim (ext2_inv_unit_sn … H) -H #H destruct +elim (frees_sex_conf … Hf … H0) -Hf -H0 +/4 width=7 by sle_sex_trans, ex2_intro/ (* note: 2 ex2_intro *) +qed-. + +(* Main properties with free variables inclusion for restricted closures ****) + +theorem rex_conf: ∀R1,R2. + rex_fsge_compatible R1 → + rex_fsge_compatible R2 → + R_confluent2_rex R1 R2 R1 R2 → + ∀T. confluent2 … (rex R1 T) (rex R2 T). +#R1 #R2 #HR1 #HR2 #HR12 #T #L0 #L1 * #f1 #Hf1 #HL01 #L2 * #f #Hf #HL02 +lapply (frees_mono … Hf1 … Hf) -Hf1 #Hf12 +lapply (sex_eq_repl_back … HL01 … Hf12) -f1 #HL01 +elim (sex_conf … HL01 … HL02) /2 width=3 by ex2_intro/ [ | -HL01 -HL02 ] +[ #L #HL1 #HL2 + elim (frees_sex_conf … Hf … HL01) // -HR1 -HL01 #f1 #Hf1 #H1 + elim (frees_sex_conf … Hf … HL02) // -HR2 -HL02 #f2 #Hf2 #H2 + lapply (sle_sex_trans … HL1 … H1) // -HL1 -H1 #HL1 + lapply (sle_sex_trans … HL2 … H2) // -HL2 -H2 #HL2 + /3 width=5 by ex2_intro/ +| #g * #I0 [2: #V0 ] #K0 #n #HLK0 #Hgf #Z1 #H1 #Z2 #H2 #K1 #HK01 #K2 #HK02 + [ elim (ext2_inv_pair_sn … H1) -H1 #V1 #HV01 #H destruct + elim (ext2_inv_pair_sn … H2) -H2 #V2 #HV02 #H destruct + elim (frees_inv_drops_next … Hf … HLK0 … Hgf) -Hf -HLK0 -Hgf #g0 #Hg0 #H0 + lapply (sle_sex_trans … HK01 … H0) // -HK01 #HK01 + lapply (sle_sex_trans … HK02 … H0) // -HK02 #HK02 + elim (HR12 … HV01 … HV02 K1 … K2) /3 width=3 by ext2_pair, ex2_intro/ + | lapply (ext2_inv_unit_sn … H1) -H1 #H destruct + lapply (ext2_inv_unit_sn … H2) -H2 #H destruct + /3 width=3 by ext2_unit, ex2_intro/ + ] +] +qed-. + +theorem rex_trans_fsle: ∀R1,R2,R3. + rex_fsle_compatible R1 → f_transitive_next R1 R2 R3 → + ∀L1,L,T. L1 ⪤[R1, T] L → + ∀L2. L ⪤[R2, T] L2 → L1 ⪤[R3, T] L2. +#R1 #R2 #R3 #H1R #H2R #L1 #L #T #H +lapply (H1R … H) -H1R #H0 +cases H -H #f1 #Hf1 #HL1 #L2 * #f2 #Hf2 #HL2 +lapply (fsle_inv_frees_eq … H0 … Hf1 … Hf2) -H0 -Hf2 +/4 width=14 by sex_trans_gen, sex_fwd_length, sle_sex_trans, ex2_intro/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/rex_length.ma b/matita/matita/contribs/lambdadelta/basic_2/static/rex_length.ma new file mode 100644 index 000000000..bd6b0d6c2 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/rex_length.ma @@ -0,0 +1,72 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/sex_length.ma". +include "basic_2/static/rex_drops.ma". + +(* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****) + +(* Forward lemmas with length for local environments ************************) + +(* Basic_2A1: uses: llpx_sn_fwd_length *) +lemma rex_fwd_length (R): ∀L1,L2,T. L1 ⪤[R, T] L2 → |L1| = |L2|. +#R #L1 #L2 #T * /2 width=4 by sex_fwd_length/ +qed-. + +(* Properties with length for local environments ****************************) + +(* Basic_2A1: uses: llpx_sn_sort *) +lemma rex_sort_length (R): ∀L1,L2. |L1| = |L2| → ∀s. L1 ⪤[R, ⋆s] L2. +#R #L1 elim L1 -L1 +[ #Y #H #s >(length_inv_zero_sn … H) -H // +| #K1 #I1 #IH #Y #H #s + elim (length_inv_succ_sn … H) -H #I2 #K2 #HK12 #H destruct + /3 width=1 by rex_sort/ +] +qed. + +(* Basic_2A1: uses: llpx_sn_gref *) +lemma rex_gref_length (R): ∀L1,L2. |L1| = |L2| → ∀l. L1 ⪤[R, §l] L2. +#R #L1 elim L1 -L1 +[ #Y #H #s >(length_inv_zero_sn … H) -H // +| #K1 #I1 #IH #Y #H #s + elim (length_inv_succ_sn … H) -H #I2 #K2 #HK12 #H destruct + /3 width=1 by rex_gref/ +] +qed. + +lemma rex_unit_length (R): ∀L1,L2. |L1| = |L2| → ∀I. L1.ⓤ{I} ⪤[R, #0] L2.ⓤ{I}. +/3 width=3 by rex_unit, sex_length_isid/ qed. + +(* Basic_2A1: uses: llpx_sn_lift_le llpx_sn_lift_ge *) +lemma rex_lifts_bi (R): d_liftable2_sn … lifts R → + ∀L1,L2. |L1| = |L2| → ∀K1,K2,T. K1 ⪤[R, T] K2 → + ∀b,f. ⬇*[b, f] L1 ≘ K1 → ⬇*[b, f] L2 ≘ K2 → + ∀U. ⬆*[f] T ≘ U → L1 ⪤[R, U] L2. +#R #HR #L1 #L2 #HL12 #K1 #K2 #T * #f1 #Hf1 #HK12 #b #f #HLK1 #HLK2 #U #HTU +elim (frees_total L1 U) #f2 #Hf2 +lapply (frees_fwd_coafter … Hf2 … HLK1 … HTU … Hf1) -HTU #Hf +/4 width=12 by sex_length_cfull, sex_liftable_co_dedropable_bi, cext2_d_liftable2_sn, cfull_lift_sn, ex2_intro/ +qed-. + +(* Inversion lemmas with length for local environment ***********************) + +lemma rex_inv_zero_length (R): ∀Y1,Y2. Y1 ⪤[R, #0] Y2 → + ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆ + | ∃∃I,L1,L2,V1,V2. L1 ⪤[R, V1] L2 & R L1 V1 V2 & + Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2 + | ∃∃I,L1,L2. |L1| = |L2| & Y1 = L1.ⓤ{I} & Y2 = L2.ⓤ{I}. +#R #Y1 #Y2 #H elim (rex_inv_zero … H) -H * +/4 width=9 by sex_fwd_length, ex4_5_intro, ex3_3_intro, or3_intro2, or3_intro1, or3_intro0, conj/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/rex_lex.ma b/matita/matita/contribs/lambdadelta/basic_2/static/rex_lex.ma new file mode 100644 index 000000000..970588d1a --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/rex_lex.ma @@ -0,0 +1,41 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/lex.ma". +include "basic_2/static/rex_fsle.ma". +include "basic_2/static/req.ma". + +(* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****) + +(* Properties with generic extension of a context-sensitive relation ********) + +lemma rex_lex: ∀R,L1,L2. L1 ⪤[R] L2 → ∀T. L1 ⪤[R, T] L2. +#R #L1 #L2 * #f #Hf #HL12 #T +elim (frees_total L1 T) #g #Hg +/4 width=5 by sex_sdj, sdj_isid_sn, ex2_intro/ +qed. + +(* Inversion lemmas with generic extension of a context sensitive relation **) + +lemma rex_inv_lex_req: ∀R. c_reflexive … R → + rex_fsge_compatible R → + ∀L1,L2,T. L1 ⪤[R, T] L2 → + ∃∃L. L1 ⪤[R] L & L ≡[T] L2. +#R #H1R #H2R #L1 #L2 #T * #f1 #Hf1 #HL +elim (sex_sdj_split … ceq_ext … HL 𝐈𝐝 ?) -HL +[ #L0 #HL10 #HL02 |*: /2 width=1 by ext2_refl, sdj_isid_dx/ ] -H1R +lapply (sex_sdj … HL10 f1 ?) /2 width=1 by sdj_isid_sn/ #H +elim (frees_sex_conf … Hf1 … H) // -H2R -H #f0 #Hf0 #Hf01 +/4 width=7 by sle_sex_trans, (* 2x *) ex2_intro/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/rex_rex.ma b/matita/matita/contribs/lambdadelta/basic_2/static/rex_rex.ma new file mode 100644 index 000000000..96162fd38 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/rex_rex.ma @@ -0,0 +1,89 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/sex_sex.ma". +include "basic_2/static/frees_fqup.ma". +include "basic_2/static/rex.ma". + +(* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****) + +(* Advanced inversion lemmas ************************************************) + +lemma rex_inv_frees: ∀R,L1,L2,T. L1 ⪤[R, T] L2 → + ∀f. L1 ⊢ 𝐅*⦃T⦄ ≘ f → L1 ⪤[cext2 R, cfull, f] L2. +#R #L1 #L2 #T * /3 width=6 by frees_mono, sex_eq_repl_back/ +qed-. + +(* Advanced properties ******************************************************) + +(* Basic_2A1: uses: llpx_sn_dec *) +lemma rex_dec: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) → + ∀L1,L2,T. Decidable (L1 ⪤[R, T] L2). +#R #HR #L1 #L2 #T +elim (frees_total L1 T) #f #Hf +elim (sex_dec (cext2 R) cfull … L1 L2 f) +/4 width=3 by rex_inv_frees, cfull_dec, ext2_dec, ex2_intro, or_intror, or_introl/ +qed-. + +(* Main properties **********************************************************) + +(* Basic_2A1: uses: llpx_sn_bind llpx_sn_bind_O *) +theorem rex_bind: ∀R,p,I,L1,L2,V1,V2,T. + L1 ⪤[R, V1] L2 → L1.ⓑ{I}V1 ⪤[R, T] L2.ⓑ{I}V2 → + L1 ⪤[R, ⓑ{p,I}V1.T] L2. +#R #p #I #L1 #L2 #V1 #V2 #T * #f1 #HV #Hf1 * #f2 #HT #Hf2 +lapply (sex_fwd_bind … Hf2) -Hf2 #Hf2 elim (sor_isfin_ex f1 (⫱f2)) +/3 width=7 by frees_fwd_isfin, frees_bind, sex_join, isfin_tl, ex2_intro/ +qed. + +(* Basic_2A1: llpx_sn_flat *) +theorem rex_flat: ∀R,I,L1,L2,V,T. + L1 ⪤[R, V] L2 → L1 ⪤[R, T] L2 → + L1 ⪤[R, ⓕ{I}V.T] L2. +#R #I #L1 #L2 #V #T * #f1 #HV #Hf1 * #f2 #HT #Hf2 elim (sor_isfin_ex f1 f2) +/3 width=7 by frees_fwd_isfin, frees_flat, sex_join, ex2_intro/ +qed. + +theorem rex_bind_void: ∀R,p,I,L1,L2,V,T. + L1 ⪤[R, V] L2 → L1.ⓧ ⪤[R, T] L2.ⓧ → + L1 ⪤[R, ⓑ{p,I}V.T] L2. +#R #p #I #L1 #L2 #V #T * #f1 #HV #Hf1 * #f2 #HT #Hf2 +lapply (sex_fwd_bind … Hf2) -Hf2 #Hf2 elim (sor_isfin_ex f1 (⫱f2)) +/3 width=7 by frees_fwd_isfin, frees_bind_void, sex_join, isfin_tl, ex2_intro/ +qed. + +(* Negated inversion lemmas *************************************************) + +(* Basic_2A1: uses: nllpx_sn_inv_bind nllpx_sn_inv_bind_O *) +lemma rnex_inv_bind: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) → + ∀p,I,L1,L2,V,T. (L1 ⪤[R, ⓑ{p,I}V.T] L2 → ⊥) → + (L1 ⪤[R, V] L2 → ⊥) ∨ (L1.ⓑ{I}V ⪤[R, T] L2.ⓑ{I}V → ⊥). +#R #HR #p #I #L1 #L2 #V #T #H elim (rex_dec … HR L1 L2 V) +/4 width=2 by rex_bind, or_intror, or_introl/ +qed-. + +(* Basic_2A1: uses: nllpx_sn_inv_flat *) +lemma rnex_inv_flat: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) → + ∀I,L1,L2,V,T. (L1 ⪤[R, ⓕ{I}V.T] L2 → ⊥) → + (L1 ⪤[R, V] L2 → ⊥) ∨ (L1 ⪤[R, T] L2 → ⊥). +#R #HR #I #L1 #L2 #V #T #H elim (rex_dec … HR L1 L2 V) +/4 width=1 by rex_flat, or_intror, or_introl/ +qed-. + +lemma rnex_inv_bind_void: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) → + ∀p,I,L1,L2,V,T. (L1 ⪤[R, ⓑ{p,I}V.T] L2 → ⊥) → + (L1 ⪤[R, V] L2 → ⊥) ∨ (L1.ⓧ ⪤[R, T] L2.ⓧ → ⊥). +#R #HR #p #I #L1 #L2 #V #T #H elim (rex_dec … HR L1 L2 V) +/4 width=2 by rex_bind_void, or_intror, or_introl/ +qed-. 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 e0976dd36..7d94ab507 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 @@ -71,7 +71,7 @@ table { } ] [ { "unbound context-sensitive parallel rst-computation" * } { - [ [ "strongly normalizing for closures" ] "fsb" + "( ≥[?,?] 𝐒⦃?,?,?⦄ )" "fsb_ffdeq" + "fsb_aaa" + "fsb_csx" + "fsb_fpbg" * ] + [ [ "strongly normalizing for closures" ] "fsb" + "( ≥[?,?] 𝐒⦃?,?,?⦄ )" "fsb_fdeq" + "fsb_aaa" + "fsb_csx" + "fsb_fpbg" * ] [ [ "proper for closures" ] "fpbg" + "( ⦃?,?,?⦄ >[?,?] ⦃?,?,?⦄ )" "fpbg_fqup" + "fpbg_cpxs" + "fpbg_lpxs" + "fpbg_fpbs" + "fpbg_fpbg" * ] [ [ "for closures" ] "fpbs" + "( ⦃?,?,?⦄ ≥[?,?] ⦃?,?,?⦄ )" "fpbs_fqup" + "fpbs_fqus" + "fpbs_aaa" + "fpbs_cpx" + "fpbs_fpb" + "fpbs_cpxs" + "fpbs_lpxs" + "fpbs_csx" + "fpbs_fpbs" * ] } @@ -80,10 +80,10 @@ table { [ [ "refinement for lenvs on selected entries" ] "lsubsx" + "( ? ⊢ ? ⊆ⓧ[?,?,?] ? )" "lsubsx_lfsx" + "lsubsx_lsubsx" * ] [ [ "strongly normalizing for lenvs on referred entries" ] "rdsx" + "( ? ⊢ ⬈*[?,?,?] 𝐒⦃?⦄ )" "rdsx_length" + "rdsx_drops" + "rdsx_fqup" + "rdsx_cpxs" + "rdsx_csx" + "rdsx_rdsx" * ] [ [ "strongly normalizing for term vectors" ] "csx_vector" + "( ⦃?,?⦄ ⊢ ⬈*[?,?] 𝐒⦃?⦄ )" "csx_cnx_vector" + "csx_csx_vector" * ] - [ [ "strongly normalizing for terms" ] "csx" + "( ⦃?,?⦄ ⊢ ⬈*[?,?] 𝐒⦃?⦄ )" "csx_simple" + "csx_simple_theq" + "csx_drops" + "csx_fqus" + "csx_lsubr" + "csx_lfdeq" + "csx_ffdeq" + "csx_aaa" + "csx_gcp" + "csx_gcr" + "csx_lpx" + "csx_cnx" + "csx_fpbq" + "csx_cpxs" + "csx_lpxs" + "csx_csx" * ] - [ [ "for lenvs on all entries" ] "lpxs" + "( ⦃?,?⦄ ⊢ ⬈*[?] ? )" "lpxs_length" + "lpxs_drops" + "lpxs_lfdeq" + "lpxs_ffdeq" + "lpxs_aaa" + "lpxs_lpx" + "lpxs_cpxs" + "lpxs_lpxs" * ] + [ [ "strongly normalizing for terms" ] "csx" + "( ⦃?,?⦄ ⊢ ⬈*[?,?] 𝐒⦃?⦄ )" "csx_simple" + "csx_simple_theq" + "csx_drops" + "csx_fqus" + "csx_lsubr" + "csx_rdeq" + "csx_fdeq" + "csx_aaa" + "csx_gcp" + "csx_gcr" + "csx_lpx" + "csx_cnx" + "csx_fpbq" + "csx_cpxs" + "csx_lpxs" + "csx_csx" * ] + [ [ "for lenvs on all entries" ] "lpxs" + "( ⦃?,?⦄ ⊢ ⬈*[?] ? )" "lpxs_length" + "lpxs_drops" + "lpxs_rdeq" + "lpxs_fdeq" + "lpxs_aaa" + "lpxs_lpx" + "lpxs_cpxs" + "lpxs_lpxs" * ] [ [ "for binders" ] "cpxs_ext" + "( ⦃?,?⦄ ⊢ ? ⬈*[?] ? )" * ] - [ [ "for terms" ] "cpxs" + "( ⦃?,?⦄ ⊢ ? ⬈*[?] ? )" "cpxs_tdeq" + "cpxs_theq" + "cpxs_theq_vector" + "cpxs_drops" + "cpxs_fqus" + "cpxs_lsubr" + "cpxs_lfdeq" + "cpxs_ffdeq" + "cpxs_aaa" + "cpxs_lpx" + "cpxs_cnx" + "cpxs_cpxs" * ] + [ [ "for terms" ] "cpxs" + "( ⦃?,?⦄ ⊢ ? ⬈*[?] ? )" "cpxs_tdeq" + "cpxs_theq" + "cpxs_theq_vector" + "cpxs_drops" + "cpxs_fqus" + "cpxs_lsubr" + "cpxs_rdeq" + "cpxs_fdeq" + "cpxs_aaa" + "cpxs_lpx" + "cpxs_cnx" + "cpxs_cpxs" * ] } ] } @@ -92,7 +92,7 @@ table { [ { "rt-transition" * } { [ { "unbound parallel rst-transition" * } { [ [ "for closures" ] "fpbq" + "( ⦃?,?,?⦄ ≽[?,?] ⦃?,?,?⦄ )" "fpbq_aaa" + "fpbq_fpb" * ] - [ [ "proper for closures" ] "fpb" + "( ⦃?,?,?⦄ ≻[?,?] ⦃?,?,?⦄ )" "fpb_lfdeq" + "fpb_ffdeq" * ] + [ [ "proper for closures" ] "fpb" + "( ⦃?,?,?⦄ ≻[?,?] ⦃?,?,?⦄ )" "fpb_rdeq" + "fpb_fdeq" * ] } ] [ { "context-sensitive parallel r-transition" * } { @@ -107,10 +107,10 @@ table { ] [ { "unbound context-sensitive parallel rt-transition" * } { [ [ "normal form for terms" ] "cnx" + "( ⦃?,?⦄ ⊢ ⬈[?,?] 𝐍⦃?⦄ )" "cnx_simple" + "cnx_drops" + "cnx_cnx" * ] - [ [ "for lenvs on referred entries" ] "lfpx" + "( ⦃?,?⦄ ⊢ ⬈[?,?] ? )" "lfpx_length" + "lfpx_fqup" + "lfpx_fsle" + "lfpx_lfdeq" + "lfpx_lpx" + "lfpx_lfpx" * ] - [ [ "for lenvs on all entries" ] "lpx" + "( ⦃?,?⦄ ⊢ ⬈[?] ? )" "lpx_length" + "lpx_drops" + "lpx_fquq" + "lpx_fsle" + "lpx_lfdeq" + "lpx_aaa" * ] + [ [ "for lenvs on referred entries" ] "rpx" + "( ⦃?,?⦄ ⊢ ⬈[?,?] ? )" "rpx_length" + "rpx_fqup" + "rpx_fsle" + "rpx_rdeq" + "rpx_lpx" + "rpx_rpx" * ] + [ [ "for lenvs on all entries" ] "lpx" + "( ⦃?,?⦄ ⊢ ⬈[?] ? )" "lpx_length" + "lpx_drops" + "lpx_fquq" + "lpx_fsle" + "lpx_rdeq" + "lpx_aaa" * ] [ [ "for binders" ] "cpx_ext" + "( ⦃?,?⦄ ⊢ ? ⬈[?] ? )" * ] - [ [ "for terms" ] "cpx" + "( ⦃?,?⦄ ⊢ ? ⬈[?] ? )" "cpx_simple" + "cpx_drops" + "cpx_fqus" + "cpx_lsubr" + "cpx_lfeq" + "cpx_lfdeq" + "cpx_ffdeq" * ] + [ [ "for terms" ] "cpx" + "( ⦃?,?⦄ ⊢ ? ⬈[?] ? )" "cpx_simple" + "cpx_drops" + "cpx_fqus" + "cpx_lsubr" + "cpx_req" + "cpx_rdeq" + "cpx_fdeq" * ] } ] [ { "bound context-sensitive parallel rt-transition" * } { @@ -122,7 +122,7 @@ table { class "water" [ { "iterated static typing" * } { [ { "iterated generic extension of a context-sensitive relation" * } { - [ [ "for lenvs on referred entries" ] "tc_lfxs" + "( ? ⦻**[?,?] ? )" "tc_lfxs_length" + "tc_lfxs_lex" + "tc_lfxs_drops" + "tc_lfxs_fqup" + "tc_lfxs_tc_lfxs" * ] + [ [ "for lenvs on referred entries" ] "rexs" + "( ? ⪤*[?,?] ? )" "rexs_length" + "rexs_lex" + "rexs_drops" + "rexs_fqup" + "rexs_rexs" * ] } ] } @@ -137,20 +137,20 @@ table { ] [ { "atomic arity assignment" * } { [ [ "restricted refinement for lenvs" ] "lsuba" + "( ? ⊢ ? ⫃⁝ ? )" "lsuba_drops" + "lsuba_lsubr" + "lsuba_aaa" + "lsuba_lsuba" * ] - [ [ "for terms" ] "aaa" + "( ⦃?,?⦄ ⊢ ? ⁝ ? )" "aaa_drops" + "aaa_fqus" + "aaa_lfdeq" + "aaa_ffdeq" + "aaa_aaa" * ] + [ [ "for terms" ] "aaa" + "( ⦃?,?⦄ ⊢ ? ⁝ ? )" "aaa_drops" + "aaa_fqus" + "aaa_rdeq" + "aaa_fdeq" + "aaa_aaa" * ] } ] [ { "degree-based equivalence" * } { - [ [ "for closures on referred entries" ] "ffdeq" + "( ⦃?,?,?⦄ ≛[?,?] ⦃?,?,?⦄ )" "ffdeq_fqup" + "ffdeq_fqus" + "ffdeq_lfeq" + "ffdeq_ffdeq" * ] - [ [ "for lenvs on referred entries" ] "lfdeq" + "( ? ≛[?,?,?] ? )" "lfdeq_length" + "lfdeq_drops" + "lfdeq_fqup" + "lfdeq_fqus" + "lfdeq_lfeq" + "lfdeq_lfdeq" * ] + [ [ "for closures on referred entries" ] "fdeq" + "( ⦃?,?,?⦄ ≛[?,?] ⦃?,?,?⦄ )" "fdeq_fqup" + "fdeq_fqus" + "fdeq_req" + "fdeq_fdeq" * ] + [ [ "for lenvs on referred entries" ] "rdeq" + "( ? ≛[?,?,?] ? )" "rdeq_length" + "rdeq_drops" + "rdeq_fqup" + "rdeq_fqus" + "rdeq_req" + "rdeq_rdeq" * ] } ] [ { "syntactic equivalence" * } { - [ [ "for lenvs on referred entries" ] "lfeq" + "( ? ≡[?] ? )" "lfeq_fqup" + "lfeq_fsle" * ] + [ [ "for lenvs on referred entries" ] "req" + "( ? ≡[?] ? )" "req_fqup" + "req_fsle" * ] } ] [ { "generic extension of a context-sensitive relation" * } { - [ [ "for lenvs on referred entries" ] "lfxs" + "( ? ⦻*[?,?] ? )" "lfxs_length" + "lfxs_lex" + "lfxs_drops" + "lfxs_fqup" + "lfxs_fsle" + "lfxs_lfxs" * ] + [ [ "for lenvs on referred entries" ] "rex" + "( ? ⪤[?,?] ? )" "rex_length" + "rex_lex" + "rex_drops" + "rex_fqup" + "rex_fsle" + "rex_rex" * ] } ] [ { "context-sensitive free variables" * } { @@ -186,7 +186,7 @@ table { class "orange" [ { "relocation" * } { [ { "generic slicing" * } { - [ [ "for lenvs" ] "drops" + "( ⬇*[?,?] ? ≘ ? )" + "( ⬇*[?] ? ≘ ? )" "drops_ctc" + "drops_ltc" + "drops_weight" + "drops_length" + "drops_cext2" + "drops_lexs" + "drops_lex" + "drops_lreq" + "drops_drops" + "drops_vector" * ] + [ [ "for lenvs" ] "drops" + "( ⬇*[?,?] ? ≘ ? )" + "( ⬇*[?] ? ≘ ? )" "drops_ctc" + "drops_ltc" + "drops_weight" + "drops_length" + "drops_cext2" + "drops_sex" + "drops_lex" + "drops_seq" + "drops_drops" + "drops_vector" * ] } ] [ { "generic relocation" * } { @@ -196,12 +196,12 @@ table { } ] [ { "syntactic equivalence" * } { - [ [ "for lenvs on selected entries" ] "lreq" + "( ? ≡[?] ? )" "lreq_length" + "lreq_lreq" * ] + [ [ "for lenvs on selected entries" ] "seq" + "( ? ≡[?] ? )" "seq_length" + "seq_seq" * ] } ] [ { "generic entrywise extension" * } { - [ [ "for lenvs of one contex-sensitive relation" ] "lex" + "( ? ⦻[?] ? )" "lex_tc" + "lex_length" + "lex_lex" * ] - [ [ "for lenvs of two contex-sensitive relations" ] "lexs" + "( ? ⦻*[?,?,?] ? )" "lexs_tc" + "lexs_length" + "lexs_lexs" * ] + [ [ "for lenvs of one contex-sensitive relation" ] "lex" + "( ? ⪤[?] ? )" "lex_tc" + "lex_length" + "lex_lex" * ] + [ [ "for lenvs of two contex-sensitive relations" ] "sex" + "( ? ⪤[?,?,?] ? )" "sex_tc" + "sex_length" + "sex_sex" * ] } ] } @@ -280,9 +280,9 @@ class "italic" { 2 } [ [ "" ] "scpes ( ⦃?,?⦄ ⊢ ? •*⬌*[?,?,?,?] ? )" "scpes_aaa" + "scpes_cpcs" + "scpes_scpes" * ] } ] - [ [ "for lenvs on referred entries" ] "lfpxs" + "( ⦃?,?⦄ ⊢ ⬈*[?,?] ? )" "lfpxs_length" + "lfpxs_drops" + "lfpxs_fqup" + "lfpxs_lfdeq" + "lfpxs_ffdeq" + "lfpxs_aaa" + "lfpxs_cpxs" + "lfpxs_lpxs" + "lfpxs_lfpxs" * ] + [ [ "for lenvs on referred entries" ] "rpxs" + "( ⦃?,?⦄ ⊢ ⬈*[?,?] ? )" "rpxs_length" + "rpxs_drops" + "rpxs_fqup" + "rpxs_rdeq" + "rpxs_fdeq" + "rpxs_aaa" + "rpxs_cpxs" + "rpxs_lpxs" + "rpxs_rpxs" * ] [ [ "for lenvs on referred entries" ] - "lfpr" + "( ⦃?,?⦄ ⊢ ➡[?,?] ? )" "lfpr_length" + "lfpr_drops" + "lfpr_fquq" + "lfpr_fqup" + "lfpr_aaa" + "lfpr_lfpx" + "lfpr_lfpr" * ] + "lfpr" + "( ⦃?,?⦄ ⊢ ➡[?,?] ? )" "lfpr_length" + "lfpr_drops" + "lfpr_fquq" + "lfpr_fqup" + "lfpr_aaa" + "lfpr_rpx" + "lfpr_lfpr" * ] [ { "evaluation for context-sensitive rt-reduction" * } { [ [ "" ] "cpxe ( ⦃?,?⦄ ⊢ ➡*[?,?] 𝐍⦃?⦄ )" * ] } diff --git a/matita/matita/contribs/lambdadelta/restore.sh b/matita/matita/contribs/lambdadelta/restore.sh index af7b5e096..6e0da30fc 100644 --- a/matita/matita/contribs/lambdadelta/restore.sh +++ b/matita/matita/contribs/lambdadelta/restore.sh @@ -1,8 +1,8 @@ #!/bin/sh -for MA in `find -name "*.ma"`; do - if [ -s ${MA}.old ]; - then echo ${MA}; mv -f ${MA}.old ${MA}; +for SRC in `find -name "*.ma" -or -name "*.tbl"`; do + if [ -s ${SRC}.old ]; + then echo ${SRC}; mv -f ${SRC}.old ${SRC}; fi done -unset MA +unset SRC