From: Ferruccio Guidi Date: Fri, 15 Nov 2019 15:22:29 +0000 (+0100) Subject: update in static_2 and basic_2 for the article X-Git-Tag: make_still_working~220 X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=commitdiff_plain;h=f677b4ef7fa20f1ab36c5ee59598865d5c1b719b update in static_2 and basic_2 for the article + commented theorems activated or parked + some renaming to match the article --- diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cnuw_simple.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cnuw_simple.etc new file mode 100644 index 000000000..5b63e8d6d --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cnuw_simple.etc @@ -0,0 +1,7 @@ +(* Advanced forward lemma with with simple terms ****************************) +(* +lemma cnuw_fwd_appl_simple (h) (G) (L): + ∀V,T. ⦃G,L⦄ ⊢ ➡𝐍𝐖*[h] ⓐV.T → 𝐒⦃T⦄. +#h #G #L #V #T #HT +elim (simple_dec_ex T) [ // ] * #p #I #W #U #H destruct +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpx_req.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpx_req.etc new file mode 100644 index 000000000..8b0fdeb69 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpx_req.etc @@ -0,0 +1,13 @@ +(* req_req is missing, with req_sym *) + +(* 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_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/etc/ntas_nta.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/ntas_nta.etc new file mode 100644 index 000000000..3bfa93807 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/ntas_nta.etc @@ -0,0 +1,11 @@ +lapply (nta_ntas … H) -H #H +elim (ntas_inv_appl_sn … H) -H * #n #p #W #U #U0 #Hn #Ha #HVW #HTU #HU #HUX #HX +[ elim (eq_or_gt n) #H destruct + [ (lifts_inv_sort1 … H) -H + /2 width=3 by teqo_sort, ex2_intro/ +| #i #f #X #H + elim (lifts_inv_lref1 … H) -H #j #Hj #H destruct + /3 width=3 by teqo_lref, lifts_lref, ex2_intro/ +| #l #f #X #H + >(lifts_inv_gref1 … H) -H + /2 width=3 by teqo_gref, ex2_intro/ +| #p #I #V1 #V2 #T1 #T2 #f #X #H + elim (lifts_inv_bind1 … H) -H #W1 #U1 #_ #_ #H destruct -V1 -T1 + elim (lifts_total V2 f) #W2 #HVW2 + elim (lifts_total T2 (⫯f)) #U2 #HTU2 + /3 width=3 by teqo_pair, lifts_bind, ex2_intro/ +| #I #V1 #V2 #T1 #T2 #f #X #H + elim (lifts_inv_flat1 … H) -H #W1 #U1 #_ #_ #H destruct -V1 -T1 + elim (lifts_total V2 f) #W2 #HVW2 + elim (lifts_total T2 f) #U2 #HTU2 + /3 width=3 by teqo_pair, lifts_flat, ex2_intro/ +] +qed-. + +lemma teqo_lifts_dx: liftable2_dx teqo. +/3 width=3 by teqo_lifts_sn, liftable2_sn_dx, teqo_sym/ qed-. + +lemma teqo_lifts_bi: liftable2_bi teqo. +/3 width=6 by teqo_lifts_sn, liftable2_sn_bi/ qed-. + +(* Inversion lemmas with sort-irrelevant outer equivalence for terms ********) + +lemma teqo_inv_lifts_bi: deliftable2_bi teqo. +#U1 #U2 #H elim H -U1 -U2 [||| * ] +[ #s1 #s2 #f #X1 #H1 #X2 #H2 + >(lifts_inv_sort2 … H1) -H1 >(lifts_inv_sort2 … H2) -H2 + /1 width=1 by teqo_sort/ +| #j #f #X1 #H1 #X2 #H2 + elim (lifts_inv_lref2 … H1) -H1 #i1 #Hj1 #H destruct + elim (lifts_inv_lref2 … H2) -H2 #i2 #Hj2 #H destruct + <(at_inj … Hj2 … Hj1) -j -f -i1 + /1 width=1 by teqo_lref/ +| #l #f #X1 #H1 #X2 #H2 + >(lifts_inv_gref2 … H1) -H1 >(lifts_inv_gref2 … H2) -H2 + /1 width=1 by teqo_gref/ +| #p #I #W1 #W2 #U1 #U2 #f #X1 #H1 #X2 #H2 + elim (lifts_inv_bind2 … H1) -H1 #V1 #T1 #_ #_ #H destruct -W1 -U1 + elim (lifts_inv_bind2 … H2) -H2 #V2 #T2 #_ #_ #H destruct -W2 -U2 + /1 width=1 by teqo_pair/ +| #I #W1 #W2 #U1 #U2 #f #X1 #H1 #X2 #H2 + elim (lifts_inv_flat2 … H1) -H1 #V1 #T1 #_ #_ #H destruct -W1 -U1 + elim (lifts_inv_flat2 … H2) -H2 #V2 #T2 #_ #_ #H destruct -W2 -U2 + /1 width=1 by teqo_pair/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/static_2/relocation/lifts_toeq.ma b/matita/matita/contribs/lambdadelta/static_2/relocation/lifts_toeq.ma deleted file mode 100644 index ffbddcbe4..000000000 --- a/matita/matita/contribs/lambdadelta/static_2/relocation/lifts_toeq.ma +++ /dev/null @@ -1,76 +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 "static_2/syntax/toeq.ma". -include "static_2/relocation/lifts_lifts.ma". - -(* GENERIC RELOCATION FOR TERMS *********************************************) - -(* Properties with sort-irrelevant outer equivalence for terms **************) - -lemma toeq_lifts_sn: liftable2_sn toeq. -#T1 #T2 #H elim H -T1 -T2 [||| * ] -[ #s1 #s2 #f #X #H - >(lifts_inv_sort1 … H) -H - /2 width=3 by toeq_sort, ex2_intro/ -| #i #f #X #H - elim (lifts_inv_lref1 … H) -H #j #Hj #H destruct - /3 width=3 by toeq_lref, lifts_lref, ex2_intro/ -| #l #f #X #H - >(lifts_inv_gref1 … H) -H - /2 width=3 by toeq_gref, ex2_intro/ -| #p #I #V1 #V2 #T1 #T2 #f #X #H - elim (lifts_inv_bind1 … H) -H #W1 #U1 #_ #_ #H destruct -V1 -T1 - elim (lifts_total V2 f) #W2 #HVW2 - elim (lifts_total T2 (⫯f)) #U2 #HTU2 - /3 width=3 by toeq_pair, lifts_bind, ex2_intro/ -| #I #V1 #V2 #T1 #T2 #f #X #H - elim (lifts_inv_flat1 … H) -H #W1 #U1 #_ #_ #H destruct -V1 -T1 - elim (lifts_total V2 f) #W2 #HVW2 - elim (lifts_total T2 f) #U2 #HTU2 - /3 width=3 by toeq_pair, lifts_flat, ex2_intro/ -] -qed-. - -lemma toeq_lifts_dx: liftable2_dx toeq. -/3 width=3 by toeq_lifts_sn, liftable2_sn_dx, toeq_sym/ qed-. - -lemma toeq_lifts_bi: liftable2_bi toeq. -/3 width=6 by toeq_lifts_sn, liftable2_sn_bi/ qed-. - -(* Inversion lemmas with sort-irrelevant outer equivalence for terms ********) - -lemma toeq_inv_lifts_bi: deliftable2_bi toeq. -#U1 #U2 #H elim H -U1 -U2 [||| * ] -[ #s1 #s2 #f #X1 #H1 #X2 #H2 - >(lifts_inv_sort2 … H1) -H1 >(lifts_inv_sort2 … H2) -H2 - /1 width=1 by toeq_sort/ -| #j #f #X1 #H1 #X2 #H2 - elim (lifts_inv_lref2 … H1) -H1 #i1 #Hj1 #H destruct - elim (lifts_inv_lref2 … H2) -H2 #i2 #Hj2 #H destruct - <(at_inj … Hj2 … Hj1) -j -f -i1 - /1 width=1 by toeq_lref/ -| #l #f #X1 #H1 #X2 #H2 - >(lifts_inv_gref2 … H1) -H1 >(lifts_inv_gref2 … H2) -H2 - /1 width=1 by toeq_gref/ -| #p #I #W1 #W2 #U1 #U2 #f #X1 #H1 #X2 #H2 - elim (lifts_inv_bind2 … H1) -H1 #V1 #T1 #_ #_ #H destruct -W1 -U1 - elim (lifts_inv_bind2 … H2) -H2 #V2 #T2 #_ #_ #H destruct -W2 -U2 - /1 width=1 by toeq_pair/ -| #I #W1 #W2 #U1 #U2 #f #X1 #H1 #X2 #H2 - elim (lifts_inv_flat2 … H1) -H1 #V1 #T1 #_ #_ #H destruct -W1 -U1 - elim (lifts_inv_flat2 … H2) -H2 #V2 #T2 #_ #_ #H destruct -W2 -U2 - /1 width=1 by toeq_pair/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/static_2/static/frees_drops.ma b/matita/matita/contribs/lambdadelta/static_2/static/frees_drops.ma index 5cbeaea87..698427fee 100644 --- a/matita/matita/contribs/lambdadelta/static_2/static/frees_drops.ma +++ b/matita/matita/contribs/lambdadelta/static_2/static/frees_drops.ma @@ -48,16 +48,7 @@ lemma frees_unit_drops: #J #L #HLK #H destruct /3 width=1 by frees_lref/ ] qed. -(* -lemma frees_sort_pushs: - ∀f,K,s. K ⊢ 𝐅+⦃⋆s⦄ ≘ f → - ∀i,L. ⇩*[i] L ≘ K → L ⊢ 𝐅+⦃⋆s⦄ ≘ ⫯*[i] f. -#f #K #s #Hf #i elim i -i -[ #L #H lapply (drops_fwd_isid … H ?) -H // -| #i #IH #L #H elim (drops_inv_succ … H) -H /3 width=1 by frees_sort/ -] -qed. -*) + lemma frees_lref_pushs: ∀f,K,j. K ⊢ 𝐅+⦃#j⦄ ≘ f → ∀i,L. ⇩*[i] L ≘ K → L ⊢ 𝐅+⦃#(i+j)⦄ ≘ ⫯*[i] f. @@ -67,16 +58,7 @@ lemma frees_lref_pushs: #I #Y #HYK #H destruct /3 width=1 by frees_lref/ ] qed. -(* -lemma frees_gref_pushs: - ∀f,K,l. K ⊢ 𝐅+⦃§l⦄ ≘ f → - ∀i,L. ⇩*[i] L ≘ K → L ⊢ 𝐅+⦃§l⦄ ≘ ⫯*[i] f. -#f #K #l #Hf #i elim i -i -[ #L #H lapply (drops_fwd_isid … H ?) -H // -| #i #IH #L #H elim (drops_inv_succ … H) -H /3 width=1 by frees_gref/ -] -qed. -*) + (* Advanced inversion lemmas ************************************************) lemma frees_inv_lref_drops: diff --git a/matita/matita/contribs/lambdadelta/static_2/static/rdeq.ma b/matita/matita/contribs/lambdadelta/static_2/static/rdeq.ma index 2e155d1ee..99fe14ee6 100644 --- a/matita/matita/contribs/lambdadelta/static_2/static/rdeq.ma +++ b/matita/matita/contribs/lambdadelta/static_2/static/rdeq.ma @@ -99,11 +99,11 @@ lemma rdeq_sort: ∀I1,I2,L1,L2,s. lemma rdeq_pair: ∀I,L1,L2,V1,V2. L1 ≛[V1] L2 → V1 ≛ V2 → L1.ⓑ{I}V1 ≛[#0] L2.ⓑ{I}V2. /2 width=1 by rex_pair/ qed. -(* -lemma rdeq_unit: ∀f,I,L1,L2. 𝐈⦃f⦄ → L1 ⪤[cdeq_ext,cfull,f] L2 → + +lemma rdeq_unit: ∀f,I,L1,L2. 𝐈⦃f⦄ → L1 ≛[f] L2 → L1.ⓤ{I} ≛[#0] L2.ⓤ{I}. /2 width=3 by rex_unit/ qed. -*) + lemma rdeq_lref: ∀I1,I2,L1,L2,i. L1 ≛[#i] L2 → L1.ⓘ{I1} ≛[#↑i] L2.ⓘ{I2}. /2 width=1 by rex_lref/ qed. @@ -125,17 +125,16 @@ lemma rdeq_inv_atom_sn: ∀Y2. ∀T:term. ⋆ ≛[T] Y2 → Y2 = ⋆. lemma rdeq_inv_atom_dx: ∀Y1. ∀T:term. Y1 ≛[T] ⋆ → Y1 = ⋆. /2 width=3 by rex_inv_atom_dx/ qed-. -(* -lemma rdeq_inv_zero: ∀Y1,Y2. Y1 ≛[#0] Y2 → - ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆ - | ∃∃I,L1,L2,V1,V2. L1 ≛[V1] L2 & V1 ≛ 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}. + +lemma rdeq_inv_zero: + ∀Y1,Y2. Y1 ≛[#0] Y2 → + ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆ + | ∃∃I,L1,L2,V1,V2. L1 ≛[V1] L2 & V1 ≛ V2 & Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2 + | ∃∃f,I,L1,L2. 𝐈⦃f⦄ & L1 ≛[f] L2 & Y1 = L1.ⓤ{I} & Y2 = L2.ⓤ{I}. #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: ∀Y1,Y2,i. Y1 ≛[#↑i] Y2 → ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆ | ∃∃I1,I2,L1,L2. L1 ≛[#i] L2 & diff --git a/matita/matita/contribs/lambdadelta/static_2/syntax/teqo.ma b/matita/matita/contribs/lambdadelta/static_2/syntax/teqo.ma new file mode 100644 index 000000000..9ac0ed76a --- /dev/null +++ b/matita/matita/contribs/lambdadelta/static_2/syntax/teqo.ma @@ -0,0 +1,159 @@ +(**************************************************************************) +(* ___ *) +(* ||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 "static_2/notation/relations/topiso_2.ma". +include "static_2/syntax/term.ma". + +(* SORT-IRRELEVANT OUTER EQUIVALENCE FOR TERMS ******************************) + +(* Basic_2A1: includes: tsts_atom tsts_pair *) +inductive teqo: relation term ≝ +| teqo_sort: ∀s1,s2. teqo (⋆s1) (⋆s2) +| teqo_lref: ∀i. teqo (#i) (#i) +| teqo_gref: ∀l. teqo (§l) (§l) +| teqo_pair: ∀I,V1,V2,T1,T2. teqo (②{I}V1.T1) (②{I}V2.T2) +. + +interpretation + "sort-irrelevant outer equivalence (term)" + 'TopIso T1 T2 = (teqo T1 T2). + +(* Basic inversion lemmas ***************************************************) + +fact teqo_inv_sort1_aux: ∀X,Y. X ⩳ Y → ∀s1. X = ⋆s1 → + ∃s2. Y = ⋆s2. +#X #Y * -X -Y +[ #s1 #s2 #s #H destruct /2 width=2 by ex_intro/ +| #i #s #H destruct +| #l #s #H destruct +| #I #V1 #V2 #T1 #T2 #s #H destruct +] +qed-. + +(* Basic_1: was just: iso_gen_sort *) +lemma teqo_inv_sort1: ∀Y,s1. ⋆s1 ⩳ Y → + ∃s2. Y = ⋆s2. +/2 width=4 by teqo_inv_sort1_aux/ qed-. + +fact teqo_inv_lref1_aux: ∀X,Y. X ⩳ Y → ∀i. X = #i → Y = #i. +#X #Y * -X -Y // +[ #s1 #s2 #j #H destruct +| #I #V1 #V2 #T1 #T2 #j #H destruct +] +qed-. + +(* Basic_1: was: iso_gen_lref *) +lemma teqo_inv_lref1: ∀Y,i. #i ⩳ Y → Y = #i. +/2 width=5 by teqo_inv_lref1_aux/ qed-. + +fact teqo_inv_gref1_aux: ∀X,Y. X ⩳ Y → ∀l. X = §l → Y = §l. +#X #Y * -X -Y // +[ #s1 #s2 #k #H destruct +| #I #V1 #V2 #T1 #T2 #k #H destruct +] +qed-. + +lemma teqo_inv_gref1: ∀Y,l. §l ⩳ Y → Y = §l. +/2 width=5 by teqo_inv_gref1_aux/ qed-. + +fact teqo_inv_pair1_aux: ∀T1,T2. T1 ⩳ T2 → + ∀J,W1,U1. T1 = ②{J}W1.U1 → + ∃∃W2,U2. T2 = ②{J}W2.U2. +#T1 #T2 * -T1 -T2 +[ #s1 #s2 #J #W1 #U1 #H destruct +| #i #J #W1 #U1 #H destruct +| #l #J #W1 #U1 #H destruct +| #I #V1 #V2 #T1 #T2 #J #W1 #U1 #H destruct /2 width=3 by ex1_2_intro/ +] +qed-. + +(* Basic_1: was: iso_gen_head *) +(* Basic_2A1: was: tsts_inv_pair1 *) +lemma teqo_inv_pair1: ∀J,W1,U1,T2. ②{J}W1.U1 ⩳ T2 → + ∃∃W2,U2. T2 = ②{J}W2. U2. +/2 width=7 by teqo_inv_pair1_aux/ qed-. + +fact teqo_inv_pair2_aux: ∀T1,T2. T1 ⩳ T2 → + ∀J,W2,U2. T2 = ②{J}W2.U2 → + ∃∃W1,U1. T1 = ②{J}W1.U1. +#T1 #T2 * -T1 -T2 +[ #s1 #s2 #J #W2 #U2 #H destruct +| #i #J #W2 #U2 #H destruct +| #l #J #W2 #U2 #H destruct +| #I #V1 #V2 #T1 #T2 #J #W2 #U2 #H destruct /2 width=3 by ex1_2_intro/ +] +qed-. + +(* Basic_2A1: was: tsts_inv_pair2 *) +lemma teqo_inv_pair2: ∀J,T1,W2,U2. T1 ⩳ ②{J}W2.U2 → + ∃∃W1,U1. T1 = ②{J}W1.U1. +/2 width=7 by teqo_inv_pair2_aux/ qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma teqo_inv_pair: ∀I1,I2,V1,V2,T1,T2. ②{I1}V1.T1 ⩳ ②{I2}V2.T2 → + I1 = I2. +#I1 #I2 #V1 #V2 #T1 #T2 #H elim (teqo_inv_pair1 … H) -H +#V0 #T0 #H destruct // +qed-. + +(* Basic properties *********************************************************) + +(* Basic_1: was: iso_refl *) +(* Basic_2A1: was: tsts_refl *) +lemma teqo_refl: reflexive … teqo. +* // +* /2 width=1 by teqo_lref, teqo_gref/ +qed. + +(* Basic_2A1: was: tsts_sym *) +lemma teqo_sym: symmetric … teqo. +#T1 #T2 * -T1 -T2 /2 width=3 by teqo_sort/ +qed-. + +(* Basic_2A1: was: tsts_dec *) +lemma teqo_dec: ∀T1,T2. Decidable (T1 ⩳ T2). +* [ * #s1 | #I1 #V1 #T1 ] * [1,3,5,7: * #s2 |*: #I2 #V2 #T2 ] +[ /3 width=1 by teqo_sort, or_introl/ +|2,3,13: + @or_intror #H + elim (teqo_inv_sort1 … H) -H #x #H destruct +|4,6,14: + @or_intror #H + lapply (teqo_inv_lref1 … H) -H #H destruct +|5: + elim (eq_nat_dec s1 s2) #Hs12 destruct /2 width=1 by or_introl/ + @or_intror #H + lapply (teqo_inv_lref1 … H) -H #H destruct /2 width=1 by/ +|7,8,15: + @or_intror #H + lapply (teqo_inv_gref1 … H) -H #H destruct +|9: + elim (eq_nat_dec s1 s2) #Hs12 destruct /2 width=1 by or_introl/ + @or_intror #H + lapply (teqo_inv_gref1 … H) -H #H destruct /2 width=1 by/ +|10,11,12: + @or_intror #H + elim (teqo_inv_pair1 … H) -H #X1 #X2 #H destruct +|16: + elim (eq_item2_dec I1 I2) #HI12 destruct + [ /3 width=1 by teqo_pair, or_introl/ ] + @or_intror #H + lapply (teqo_inv_pair … H) -H /2 width=1 by/ +] +qed-. + +(* Basic_2A1: removed theorems 2: + tsts_inv_atom1 tsts_inv_atom2 +*) diff --git a/matita/matita/contribs/lambdadelta/static_2/syntax/teqo_simple.ma b/matita/matita/contribs/lambdadelta/static_2/syntax/teqo_simple.ma new file mode 100644 index 000000000..8159a89f3 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/static_2/syntax/teqo_simple.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 "static_2/syntax/term_simple.ma". +include "static_2/syntax/teqo.ma". + +(* SORT-IRRELEVANT OUTER EQUIVALENCE FOR TERMS ******************************) + +(* Properies with simple (neutral) terms ************************************) + +(* Basic_2A1: was: simple_tsts_repl_dx *) +lemma simple_teqo_repl_dx: ∀T1,T2. T1 ⩳ T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄. +#T1 #T2 * -T1 -T2 // +#I #V1 #V2 #T1 #T2 #H +elim (simple_inv_pair … H) -H #J #H destruct // +qed-. + +(* Basic_2A1: was: simple_tsts_repl_sn *) +lemma simple_teqo_repl_sn: ∀T1,T2. T1 ⩳ T2 → 𝐒⦃T2⦄ → 𝐒⦃T1⦄. +/3 width=3 by simple_teqo_repl_dx, teqo_sym/ qed-. diff --git a/matita/matita/contribs/lambdadelta/static_2/syntax/teqo_simple_vector.ma b/matita/matita/contribs/lambdadelta/static_2/syntax/teqo_simple_vector.ma new file mode 100644 index 000000000..43ea6c691 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/static_2/syntax/teqo_simple_vector.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 "static_2/syntax/term_vector.ma". +include "static_2/syntax/teqo_simple.ma". + +(* SORT-IRRELEVANT OUTER EQUIVALENCE FOR TERMS ******************************) + +(* Advanced inversion lemmas with simple (neutral) terms ********************) + +(* Basic_1: was only: iso_flats_lref_bind_false iso_flats_flat_bind_false *) +(* Basic_2A1: was: tsts_inv_bind_applv_simple *) +lemma teqo_inv_applv_bind_simple (p) (I): + ∀Vs,V2,T1,T2. ⒶVs.T1 ⩳ ⓑ{p,I}V2.T2 → 𝐒⦃T1⦄ → ⊥. +#p #I #Vs #V2 #T1 #T2 #H elim (teqo_inv_pair2 … H) -H +#V0 #T0 elim Vs -Vs normalize +[ #H destruct #H /2 width=5 by simple_inv_bind/ +| #V #Vs #_ #H destruct +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/static_2/syntax/teqo_tdeq.ma b/matita/matita/contribs/lambdadelta/static_2/syntax/teqo_tdeq.ma new file mode 100644 index 000000000..a0a93b9bd --- /dev/null +++ b/matita/matita/contribs/lambdadelta/static_2/syntax/teqo_tdeq.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 "static_2/syntax/tdeq.ma". +include "static_2/syntax/teqo.ma". + +(* SORT-IRRELEVANT OUTER EQUIVALENCE FOR TERMS ******************************) + +(* Properties with sort-irrelevant equivalence for terms ********************) + +lemma tdeq_teqo: ∀T1,T2. T1 ≛ T2 → T1 ⩳ T2. +#T1 #T2 * -T1 -T2 /2 width=1 by teqo_sort, teqo_pair/ +qed. diff --git a/matita/matita/contribs/lambdadelta/static_2/syntax/teqo_teqo.ma b/matita/matita/contribs/lambdadelta/static_2/syntax/teqo_teqo.ma new file mode 100644 index 000000000..c3d6d0c3a --- /dev/null +++ b/matita/matita/contribs/lambdadelta/static_2/syntax/teqo_teqo.ma @@ -0,0 +1,40 @@ +(**************************************************************************) +(* ___ *) +(* ||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 "static_2/syntax/teqo.ma". + +(* SORT-IRRELEVANT OUTER EQUIVALENCE FOR TERMS ******************************) + +(* Main properties **********************************************************) + +(* Basic_1: was: iso_trans *) +(* Basic_2A1: was: tsts_trans *) +theorem teqo_trans: Transitive … teqo. +#T1 #T * -T1 -T +[ #s1 #s #X #H + elim (teqo_inv_sort1 … H) -s /2 width=1 by teqo_sort/ +| #i1 #i #H <(teqo_inv_lref1 … H) -H // +| #l1 #l #H <(teqo_inv_gref1 … H) -H // +| #I #V1 #V #T1 #T #X #H + elim (teqo_inv_pair1 … H) -H #V2 #T2 #H destruct // +] +qed-. + +(* Basic_2A1: was: tsts_canc_sn *) +theorem teqo_canc_sn: left_cancellable … teqo. +/3 width=3 by teqo_trans, teqo_sym/ qed-. + +(* Basic_2A1: was: tsts_canc_dx *) +theorem teqo_canc_dx: right_cancellable … teqo. +/3 width=3 by teqo_trans, teqo_sym/ qed-. diff --git a/matita/matita/contribs/lambdadelta/static_2/syntax/toeq.ma b/matita/matita/contribs/lambdadelta/static_2/syntax/toeq.ma deleted file mode 100644 index 7206a1d62..000000000 --- a/matita/matita/contribs/lambdadelta/static_2/syntax/toeq.ma +++ /dev/null @@ -1,159 +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 "static_2/notation/relations/topiso_2.ma". -include "static_2/syntax/term.ma". - -(* SORT-IRRELEVANT OUTER EQUIVALENCE FOR TERMS ******************************) - -(* Basic_2A1: includes: tsts_atom tsts_pair *) -inductive toeq: relation term ≝ -| toeq_sort: ∀s1,s2. toeq (⋆s1) (⋆s2) -| toeq_lref: ∀i. toeq (#i) (#i) -| toeq_gref: ∀l. toeq (§l) (§l) -| toeq_pair: ∀I,V1,V2,T1,T2. toeq (②{I}V1.T1) (②{I}V2.T2) -. - -interpretation - "sort-irrelevant outer equivalence (term)" - 'TopIso T1 T2 = (toeq T1 T2). - -(* Basic inversion lemmas ***************************************************) - -fact toeq_inv_sort1_aux: ∀X,Y. X ⩳ Y → ∀s1. X = ⋆s1 → - ∃s2. Y = ⋆s2. -#X #Y * -X -Y -[ #s1 #s2 #s #H destruct /2 width=2 by ex_intro/ -| #i #s #H destruct -| #l #s #H destruct -| #I #V1 #V2 #T1 #T2 #s #H destruct -] -qed-. - -(* Basic_1: was just: iso_gen_sort *) -lemma toeq_inv_sort1: ∀Y,s1. ⋆s1 ⩳ Y → - ∃s2. Y = ⋆s2. -/2 width=4 by toeq_inv_sort1_aux/ qed-. - -fact toeq_inv_lref1_aux: ∀X,Y. X ⩳ Y → ∀i. X = #i → Y = #i. -#X #Y * -X -Y // -[ #s1 #s2 #j #H destruct -| #I #V1 #V2 #T1 #T2 #j #H destruct -] -qed-. - -(* Basic_1: was: iso_gen_lref *) -lemma toeq_inv_lref1: ∀Y,i. #i ⩳ Y → Y = #i. -/2 width=5 by toeq_inv_lref1_aux/ qed-. - -fact toeq_inv_gref1_aux: ∀X,Y. X ⩳ Y → ∀l. X = §l → Y = §l. -#X #Y * -X -Y // -[ #s1 #s2 #k #H destruct -| #I #V1 #V2 #T1 #T2 #k #H destruct -] -qed-. - -lemma toeq_inv_gref1: ∀Y,l. §l ⩳ Y → Y = §l. -/2 width=5 by toeq_inv_gref1_aux/ qed-. - -fact toeq_inv_pair1_aux: ∀T1,T2. T1 ⩳ T2 → - ∀J,W1,U1. T1 = ②{J}W1.U1 → - ∃∃W2,U2. T2 = ②{J}W2.U2. -#T1 #T2 * -T1 -T2 -[ #s1 #s2 #J #W1 #U1 #H destruct -| #i #J #W1 #U1 #H destruct -| #l #J #W1 #U1 #H destruct -| #I #V1 #V2 #T1 #T2 #J #W1 #U1 #H destruct /2 width=3 by ex1_2_intro/ -] -qed-. - -(* Basic_1: was: iso_gen_head *) -(* Basic_2A1: was: tsts_inv_pair1 *) -lemma toeq_inv_pair1: ∀J,W1,U1,T2. ②{J}W1.U1 ⩳ T2 → - ∃∃W2,U2. T2 = ②{J}W2. U2. -/2 width=7 by toeq_inv_pair1_aux/ qed-. - -fact toeq_inv_pair2_aux: ∀T1,T2. T1 ⩳ T2 → - ∀J,W2,U2. T2 = ②{J}W2.U2 → - ∃∃W1,U1. T1 = ②{J}W1.U1. -#T1 #T2 * -T1 -T2 -[ #s1 #s2 #J #W2 #U2 #H destruct -| #i #J #W2 #U2 #H destruct -| #l #J #W2 #U2 #H destruct -| #I #V1 #V2 #T1 #T2 #J #W2 #U2 #H destruct /2 width=3 by ex1_2_intro/ -] -qed-. - -(* Basic_2A1: was: tsts_inv_pair2 *) -lemma toeq_inv_pair2: ∀J,T1,W2,U2. T1 ⩳ ②{J}W2.U2 → - ∃∃W1,U1. T1 = ②{J}W1.U1. -/2 width=7 by toeq_inv_pair2_aux/ qed-. - -(* Advanced inversion lemmas ************************************************) - -lemma toeq_inv_pair: ∀I1,I2,V1,V2,T1,T2. ②{I1}V1.T1 ⩳ ②{I2}V2.T2 → - I1 = I2. -#I1 #I2 #V1 #V2 #T1 #T2 #H elim (toeq_inv_pair1 … H) -H -#V0 #T0 #H destruct // -qed-. - -(* Basic properties *********************************************************) - -(* Basic_1: was: iso_refl *) -(* Basic_2A1: was: tsts_refl *) -lemma toeq_refl: reflexive … toeq. -* // -* /2 width=1 by toeq_lref, toeq_gref/ -qed. - -(* Basic_2A1: was: tsts_sym *) -lemma toeq_sym: symmetric … toeq. -#T1 #T2 * -T1 -T2 /2 width=3 by toeq_sort/ -qed-. - -(* Basic_2A1: was: tsts_dec *) -lemma toeq_dec: ∀T1,T2. Decidable (T1 ⩳ T2). -* [ * #s1 | #I1 #V1 #T1 ] * [1,3,5,7: * #s2 |*: #I2 #V2 #T2 ] -[ /3 width=1 by toeq_sort, or_introl/ -|2,3,13: - @or_intror #H - elim (toeq_inv_sort1 … H) -H #x #H destruct -|4,6,14: - @or_intror #H - lapply (toeq_inv_lref1 … H) -H #H destruct -|5: - elim (eq_nat_dec s1 s2) #Hs12 destruct /2 width=1 by or_introl/ - @or_intror #H - lapply (toeq_inv_lref1 … H) -H #H destruct /2 width=1 by/ -|7,8,15: - @or_intror #H - lapply (toeq_inv_gref1 … H) -H #H destruct -|9: - elim (eq_nat_dec s1 s2) #Hs12 destruct /2 width=1 by or_introl/ - @or_intror #H - lapply (toeq_inv_gref1 … H) -H #H destruct /2 width=1 by/ -|10,11,12: - @or_intror #H - elim (toeq_inv_pair1 … H) -H #X1 #X2 #H destruct -|16: - elim (eq_item2_dec I1 I2) #HI12 destruct - [ /3 width=1 by toeq_pair, or_introl/ ] - @or_intror #H - lapply (toeq_inv_pair … H) -H /2 width=1 by/ -] -qed-. - -(* Basic_2A1: removed theorems 2: - tsts_inv_atom1 tsts_inv_atom2 -*) diff --git a/matita/matita/contribs/lambdadelta/static_2/syntax/toeq_simple.ma b/matita/matita/contribs/lambdadelta/static_2/syntax/toeq_simple.ma deleted file mode 100644 index 00d96949c..000000000 --- a/matita/matita/contribs/lambdadelta/static_2/syntax/toeq_simple.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 "static_2/syntax/term_simple.ma". -include "static_2/syntax/toeq.ma". - -(* SORT-IRRELEVANT OUTER EQUIVALENCE FOR TERMS ******************************) - -(* Properies with simple (neutral) terms ************************************) - -(* Basic_2A1: was: simple_tsts_repl_dx *) -lemma simple_toeq_repl_dx: ∀T1,T2. T1 ⩳ T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄. -#T1 #T2 * -T1 -T2 // -#I #V1 #V2 #T1 #T2 #H -elim (simple_inv_pair … H) -H #J #H destruct // -qed-. - -(* Basic_2A1: was: simple_tsts_repl_sn *) -lemma simple_toeq_repl_sn: ∀T1,T2. T1 ⩳ T2 → 𝐒⦃T2⦄ → 𝐒⦃T1⦄. -/3 width=3 by simple_toeq_repl_dx, toeq_sym/ qed-. diff --git a/matita/matita/contribs/lambdadelta/static_2/syntax/toeq_simple_vector.ma b/matita/matita/contribs/lambdadelta/static_2/syntax/toeq_simple_vector.ma deleted file mode 100644 index 92e28c878..000000000 --- a/matita/matita/contribs/lambdadelta/static_2/syntax/toeq_simple_vector.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 "static_2/syntax/term_vector.ma". -include "static_2/syntax/toeq_simple.ma". - -(* SORT-IRRELEVANT OUTER EQUIVALENCE FOR TERMS ******************************) - -(* Advanced inversion lemmas with simple (neutral) terms ********************) - -(* Basic_1: was only: iso_flats_lref_bind_false iso_flats_flat_bind_false *) -(* Basic_2A1: was: tsts_inv_bind_applv_simple *) -lemma toeq_inv_applv_bind_simple (p) (I): - ∀Vs,V2,T1,T2. ⒶVs.T1 ⩳ ⓑ{p,I}V2.T2 → 𝐒⦃T1⦄ → ⊥. -#p #I #Vs #V2 #T1 #T2 #H elim (toeq_inv_pair2 … H) -H -#V0 #T0 elim Vs -Vs normalize -[ #H destruct #H /2 width=5 by simple_inv_bind/ -| #V #Vs #_ #H destruct -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/static_2/syntax/toeq_tdeq.ma b/matita/matita/contribs/lambdadelta/static_2/syntax/toeq_tdeq.ma deleted file mode 100644 index 73be8f600..000000000 --- a/matita/matita/contribs/lambdadelta/static_2/syntax/toeq_tdeq.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 "static_2/syntax/tdeq.ma". -include "static_2/syntax/toeq.ma". - -(* SORT-IRRELEVANT OUTER EQUIVALENCE FOR TERMS ******************************) - -(* Properties with sort-irrelevant equivalence for terms ********************) - -lemma tdeq_toeq: ∀T1,T2. T1 ≛ T2 → T1 ⩳ T2. -#T1 #T2 * -T1 -T2 /2 width=1 by toeq_sort, toeq_pair/ -qed. diff --git a/matita/matita/contribs/lambdadelta/static_2/syntax/toeq_toeq.ma b/matita/matita/contribs/lambdadelta/static_2/syntax/toeq_toeq.ma deleted file mode 100644 index 9b7af0a51..000000000 --- a/matita/matita/contribs/lambdadelta/static_2/syntax/toeq_toeq.ma +++ /dev/null @@ -1,40 +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 "static_2/syntax/toeq.ma". - -(* SORT-IRRELEVANT OUTER EQUIVALENCE FOR TERMS ******************************) - -(* Main properties **********************************************************) - -(* Basic_1: was: iso_trans *) -(* Basic_2A1: was: tsts_trans *) -theorem toeq_trans: Transitive … toeq. -#T1 #T * -T1 -T -[ #s1 #s #X #H - elim (toeq_inv_sort1 … H) -s /2 width=1 by toeq_sort/ -| #i1 #i #H <(toeq_inv_lref1 … H) -H // -| #l1 #l #H <(toeq_inv_gref1 … H) -H // -| #I #V1 #V #T1 #T #X #H - elim (toeq_inv_pair1 … H) -H #V2 #T2 #H destruct // -] -qed-. - -(* Basic_2A1: was: tsts_canc_sn *) -theorem toeq_canc_sn: left_cancellable … toeq. -/3 width=3 by toeq_trans, toeq_sym/ qed-. - -(* Basic_2A1: was: tsts_canc_dx *) -theorem toeq_canc_dx: right_cancellable … toeq. -/3 width=3 by toeq_trans, toeq_sym/ qed-. diff --git a/matita/matita/contribs/lambdadelta/static_2/web/static_2_src.tbl b/matita/matita/contribs/lambdadelta/static_2/web/static_2_src.tbl index 44690009a..627162218 100644 --- a/matita/matita/contribs/lambdadelta/static_2/web/static_2_src.tbl +++ b/matita/matita/contribs/lambdadelta/static_2/web/static_2_src.tbl @@ -86,7 +86,7 @@ table { [ { "generic and uniform relocation" * } { [ [ "for binders" ] "lifts_bind" + "( ⇧*[?] ? ≘ ? )" "lifts_weight_bind" + "lifts_lifts_bind" * ] [ [ "for term vectors" ] "lifts_vector" + "( ⇧*[?] ? ≘ ? )" "lifts_lifts_vector" * ] - [ [ "for terms" ] "lifts" + "( ⇧*[?] ? ≘ ? )" "lifts_simple" + "lifts_weight" + "lifts_tdeq" + "lifts_tweq" + "lifts_toeq" + "lifts_lifts" * ] + [ [ "for terms" ] "lifts" + "( ⇧*[?] ? ≘ ? )" "lifts_simple" + "lifts_weight" + "lifts_tdeq" + "lifts_tweq" + "lifts_teqo" + "lifts_lifts" * ] } ] [ { "syntactic equivalence" * } { @@ -117,7 +117,7 @@ table { } ] [ { "sort-irrelevant outer equivalence" * } { - [ [ "for terms" ] "toeq" + "( ? ⩳ ? )" "toeq_simple" + "toeq_tdeq" + "toeq_toeq" + "toeq_simple_vector" * ] + [ [ "for terms" ] "teqo" + "( ? ⩳ ? )" "teqo_simple" + "teqo_tdeq" + "teqo_teqo" + "teqo_simple_vector" * ] } ] [ { "sort-irrelevant whd equivalence" * } {