X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fsyntax%2Fvoids.ma;h=854737b4a47c8eeacc3740340ecc3784e66e0fb1;hp=ca5c4fc58a691c8537e4c7d55a9780147e41743a;hb=1c8e230b1d81491b38126900d76201fb84303ced;hpb=8653dd54c57943e28e3ef60d2d0cbc1861a76a33 diff --git a/matita/matita/contribs/lambdadelta/basic_2/syntax/voids.ma b/matita/matita/contribs/lambdadelta/basic_2/syntax/voids.ma index ca5c4fc58..854737b4a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/syntax/voids.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/syntax/voids.ma @@ -12,52 +12,80 @@ (* *) (**************************************************************************) -include "basic_2/notation/relations/rvoidstar_4.ma". +include "basic_2/notation/functions/voidstar_2.ma". include "basic_2/syntax/lenv.ma". -(* EQUIVALENCE FOR LOCAL ENVIRONMENTS UP TO EXCLUSION BINDERS ***************) +(* EXTENSION OF A LOCAL ENVIRONMENT WITH EXCLUSION BINDERS ******************) -inductive voids: bi_relation nat lenv ≝ -| voids_atom : voids 0 (⋆) 0 (⋆) -| voids_pair_sn: ∀I1,I2,K1,K2,V1,n. voids n K1 n K2 → - voids 0 (K1.ⓑ{I1}V1) 0 (K2.ⓘ{I2}) -| voids_pair_dx: ∀I1,I2,K1,K2,V2,n. voids n K1 n K2 → - voids 0 (K1.ⓘ{I1}) 0 (K2.ⓑ{I2}V2) -| voids_void_sn: ∀K1,K2,n1,n2. voids n1 K1 n2 K2 → - voids (⫯n1) (K1.ⓧ) n2 K2 -| voids_void_dx: ∀K1,K2,n1,n2. voids n1 K1 n2 K2 → - voids n1 K1 (⫯n2) (K2.ⓧ) -. +rec definition voids (L:lenv) (n:nat) on n: lenv ≝ match n with +[ O ⇒ L | S m ⇒ (voids L m).ⓧ ]. -interpretation "equivalence up to exclusion binders (local environment)" - 'RVoidStar n1 L1 n2 L2 = (voids n1 L1 n2 L2). +interpretation "extension with exclusion binders (local environment)" + 'VoidStar n L = (voids L n). (* Basic properties *********************************************************) -lemma voids_refl: ∀L. ∃n. ⓧ*[n]L ≋ ⓧ*[n]L. -#L elim L -L /2 width=2 by ex_intro, voids_atom/ -#L #I * #n #IH cases I -I /3 width=2 by ex_intro, voids_pair_dx/ -* /4 width=2 by ex_intro, voids_void_sn, voids_void_dx/ +lemma voids_zero: ∀L. L = ⓧ*[0]L. +// qed. + +lemma voids_succ: ∀L,n. (ⓧ*[n]L).ⓧ = ⓧ*[⫯n]L. +// qed. + +(* Advanced properties ******************************************************) + +lemma voids_next: ∀n,L. ⓧ*[n](L.ⓧ) = ⓧ*[⫯n]L. +#n elim n -n // +qed. + +(* Basic inversion lemmas ***************************************************) + +lemma voids_atom_inv: ∀K,n. ⓧ*[n]K = ⋆ → ∧∧ ⋆ = K & 0 = n. +#K * /2 width=1 by conj/ +#n minus_S_S #H + elim (destruct_lbind_lbind_aux … H) -H #HK #_ (**) (* destruct lemma needed *) + elim (IH … HK) -IH -HK #H #Hn destruct /3 width=1 by conj, le_S_S/ + ] ] qed-. -lemma voids_inv_void_dx: ∀L1,L2,n1,n2. ⓧ*[n1]L1 ≋ ⓧ*[⫯n2]L2.ⓧ → ⓧ*[n1]L1 ≋ ⓧ*[n2]L2. -/2 width=5 by voids_inv_void_dx_aux/ qed-. +lemma voids_inv_pair_sn: ∀I,V,n1,K1,K2,n2. ⓧ*[n1]K1.ⓑ{I}V = ⓧ*[n2]K2 → + ∧∧ ⓧ*[n1-n2]K1.ⓑ{I}V = K2 & n2 ≤ n1. +#I #V #n1 elim n1 -n1 +[ #K1 #K2 minus_S_S #H + elim (destruct_lbind_lbind_aux … H) -H #HK #_ (**) (* destruct lemma needed *) + elim (IH … HK) -IH -HK #H #Hn destruct /3 width=1 by conj, le_S_S/ + ] +] +qed-. + +(* Main inversion properties ************************************************) + +theorem voids_inj: ∀n. injective … (λL. ⓧ*[n]L). +#n elim n -n // +#n #IH #L1 #L2 +