X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fstatic%2Ffeqx_feqx.ma;h=c00a1d9df8ea581dd365a39ab4a5cc7399ed7697;hb=cc178d85bc4fec05b6a9dd176f338b3275beb3d9;hp=dddcb113b99450b11fd6871da1b90fdf95a49a19;hpb=e23331eef5817eaa6c5e1c442d1d6bbb18650573;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/static_2/static/feqx_feqx.ma b/matita/matita/contribs/lambdadelta/static_2/static/feqx_feqx.ma index dddcb113b..c00a1d9df 100644 --- a/matita/matita/contribs/lambdadelta/static_2/static/feqx_feqx.ma +++ b/matita/matita/contribs/lambdadelta/static_2/static/feqx_feqx.ma @@ -20,7 +20,7 @@ include "static_2/static/feqx.ma". (* Advanced properties ******************************************************) lemma feqx_dec (G1) (G2) (L1) (L2) (T1) (T2): - Decidable (❪G1,L1,T1❫ ≅ ❪G2,L2,T2❫). + Decidable (❨G1,L1,T1❩ ≅ ❨G2,L2,T2❩). /3 width=1 by feqg_dec, sfull_dec/ qed-. (* lemma feqx_sym: tri_symmetric … feqx. @@ -36,18 +36,18 @@ theorem feqx_trans: tri_transitive … feqx. /4 width=5 by feqx_intro_sn, reqx_trans, teqx_reqx_div, teqx_trans/ qed-. -theorem feqx_canc_sn: ∀G,G1,L,L1,T,T1. ❪G,L,T❫ ≛ ❪G1,L1,T1❫→ - ∀G2,L2,T2. ❪G,L,T❫ ≛ ❪G2,L2,T2❫ → ❪G1,L1,T1❫ ≛ ❪G2,L2,T2❫. +theorem feqx_canc_sn: ∀G,G1,L,L1,T,T1. ❨G,L,T❩ ≛ ❨G1,L1,T1❩→ + ∀G2,L2,T2. ❨G,L,T❩ ≛ ❨G2,L2,T2❩ → ❨G1,L1,T1❩ ≛ ❨G2,L2,T2❩. /3 width=5 by feqx_trans, feqx_sym/ qed-. -theorem feqx_canc_dx: ∀G1,G,L1,L,T1,T. ❪G1,L1,T1❫ ≛ ❪G,L,T❫ → - ∀G2,L2,T2. ❪G2,L2,T2❫ ≛ ❪G,L,T❫ → ❪G1,L1,T1❫ ≛ ❪G2,L2,T2❫. +theorem feqx_canc_dx: ∀G1,G,L1,L,T1,T. ❨G1,L1,T1❩ ≛ ❨G,L,T❩ → + ∀G2,L2,T2. ❨G2,L2,T2❩ ≛ ❨G,L,T❩ → ❨G1,L1,T1❩ ≛ ❨G2,L2,T2❩. /3 width=5 by feqx_trans, feqx_sym/ qed-. (* Main inversion lemmas with degree-based equivalence on terms *************) -theorem feqx_tneqx_repl_dx: ∀G1,G2,L1,L2,T1,T2. ❪G1,L1,T1❫ ≛ ❪G2,L2,T2❫ → - ∀U1,U2. ❪G1,L1,U1❫ ≛ ❪G2,L2,U2❫ → +theorem feqx_tneqx_repl_dx: ∀G1,G2,L1,L2,T1,T2. ❨G1,L1,T1❩ ≛ ❨G2,L2,T2❩ → + ∀U1,U2. ❨G1,L1,U1❩ ≛ ❨G2,L2,U2❩ → (T2 ≛ U2 → ⊥) → (T1 ≛ U1 → ⊥). #G1 #G2 #L1 #L2 #T1 #T2 #HT #U1 #U2 #HU #HnTU2 #HTU1 elim (feqx_inv_gen_sn … HT) -HT #_ #_ #HT