X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fstatic%2Freqx_reqx.ma;h=8f9999e64107a6d61d3f5e237e3072b5c2e2539c;hb=e23331eef5817eaa6c5e1c442d1d6bbb18650573;hp=f7322ee5e7247b93902e70180166b6a33b981400;hpb=c7b50fec51b9a25d5bc536f44e54179fd53efb44;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/static_2/static/reqx_reqx.ma b/matita/matita/contribs/lambdadelta/static_2/static/reqx_reqx.ma index f7322ee5e..8f9999e64 100644 --- a/matita/matita/contribs/lambdadelta/static_2/static/reqx_reqx.ma +++ b/matita/matita/contribs/lambdadelta/static_2/static/reqx_reqx.ma @@ -12,37 +12,35 @@ (* *) (**************************************************************************) -include "static_2/syntax/ext2_ext2.ma". -include "static_2/syntax/teqx_teqx.ma". -include "static_2/static/reqx_length.ma". +include "static_2/static/reqg_reqg.ma". +include "static_2/static/reqx.ma". (* SORT-IRRELEVANT EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ***) (* Advanced properties ******************************************************) - -(* Basic_2A1: uses: lleq_sym *) +(* lemma reqx_sym: ∀T. symmetric … (reqx T). /3 width=3 by reqx_fsge_comp, rex_sym, teqx_sym/ qed-. - +*) (* Basic_2A1: uses: lleq_dec *) -lemma reqx_dec: ∀L1,L2. ∀T:term. Decidable (L1 ≛[T] L2). -/3 width=1 by rex_dec, teqx_dec/ qed-. - +lemma reqx_dec: ∀L1,L2. ∀T:term. Decidable (L1 ≅[T] L2). +/3 width=1 by reqg_dec, sfull_dec/ qed-. +(* (* Main properties **********************************************************) (* Basic_2A1: uses: lleq_bind lleq_bind_O *) theorem reqx_bind: ∀p,I,L1,L2,V1,V2,T. - L1 ≛[V1] L2 → L1.ⓑ{I}V1 ≛[T] L2.ⓑ{I}V2 → - L1 ≛[ⓑ{p,I}V1.T] L2. + L1 ≛[V1] L2 → L1.ⓑ[I]V1 ≛[T] L2.ⓑ[I]V2 → + L1 ≛[ⓑ[p,I]V1.T] L2. /2 width=2 by rex_bind/ qed. (* Basic_2A1: uses: lleq_flat *) theorem reqx_flat: ∀I,L1,L2,V,T. - L1 ≛[V] L2 → L1 ≛[T] L2 → L1 ≛[ⓕ{I}V.T] L2. + L1 ≛[V] L2 → L1 ≛[T] L2 → L1 ≛[ⓕ[I]V.T] L2. /2 width=1 by rex_flat/ qed. theorem reqx_bind_void: ∀p,I,L1,L2,V,T. - L1 ≛[V] L2 → L1.ⓧ ≛[T] L2.ⓧ → L1 ≛[ⓑ{p,I}V.T] L2. + L1 ≛[V] L2 → L1.ⓧ ≛[T] L2.ⓧ → L1 ≛[ⓑ[p,I]V.T] L2. /2 width=1 by rex_bind_void/ qed. (* Basic_2A1: uses: lleq_trans *) @@ -61,9 +59,9 @@ theorem reqx_canc_sn: ∀T. left_cancellable … (reqx T). theorem reqx_canc_dx: ∀T. right_cancellable … (reqx T). /3 width=3 by reqx_trans, reqx_sym/ qed-. -theorem reqx_repl: ∀L1,L2. ∀T:term. L1 ≛[T] L2 → - ∀K1. L1 ≛[T] K1 → ∀K2. L2 ≛[T] K2 → K1 ≛[T] K2. -/3 width=3 by reqx_canc_sn, reqx_trans/ qed-. +theorem reqx_repl: ∀L1,L2. ∀T:term. L1 ≅[T] L2 → + ∀K1. L1 ≅[T] K1 → ∀K2. L2 ≅[T] K2 → K1 ≅[T] K2. +/2 width=5 by reqg_repl/ qed-. (* Negated properties *******************************************************) @@ -85,15 +83,16 @@ theorem rneqx_reqx_canc_dx: ∀L1,L. ∀T:term. (L1 ≛[T] L → ⊥) → (* Negated inversion lemmas *************************************************) (* Basic_2A1: uses: nlleq_inv_bind nlleq_inv_bind_O *) -lemma rneqx_inv_bind: ∀p,I,L1,L2,V,T. (L1 ≛[ⓑ{p,I}V.T] L2 → ⊥) → - (L1 ≛[V] L2 → ⊥) ∨ (L1.ⓑ{I}V ≛[T] L2.ⓑ{I}V → ⊥). +lemma rneqx_inv_bind: ∀p,I,L1,L2,V,T. (L1 ≛[ⓑ[p,I]V.T] L2 → ⊥) → + (L1 ≛[V] L2 → ⊥) ∨ (L1.ⓑ[I]V ≛[T] L2.ⓑ[I]V → ⊥). /3 width=2 by rnex_inv_bind, teqx_dec/ qed-. (* Basic_2A1: uses: nlleq_inv_flat *) -lemma rneqx_inv_flat: ∀I,L1,L2,V,T. (L1 ≛[ⓕ{I}V.T] L2 → ⊥) → +lemma rneqx_inv_flat: ∀I,L1,L2,V,T. (L1 ≛[ⓕ[I]V.T] L2 → ⊥) → (L1 ≛[V] L2 → ⊥) ∨ (L1 ≛[T] L2 → ⊥). /3 width=2 by rnex_inv_flat, teqx_dec/ qed-. -lemma rneqx_inv_bind_void: ∀p,I,L1,L2,V,T. (L1 ≛[ⓑ{p,I}V.T] L2 → ⊥) → +lemma rneqx_inv_bind_void: ∀p,I,L1,L2,V,T. (L1 ≛[ⓑ[p,I]V.T] L2 → ⊥) → (L1 ≛[V] L2 → ⊥) ∨ (L1.ⓧ ≛[T] L2.ⓧ → ⊥). /3 width=3 by rnex_inv_bind_void, teqx_dec/ qed-. +*)