(**************************************************************************) (* ___ *) (* ||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/ynat/ynat_lt.ma". include "basic_2/notation/relations/coeq_4.ma". include "basic_2/grammar/lenv_length.ma". (* COEQUIVALENCE FOR LOCAL ENVIRONMENTS *************************************) inductive lcoeq: relation4 ynat ynat lenv lenv ≝ | lcoeq_atom: ∀d,e. lcoeq d e (⋆) (⋆) | lcoeq_zero: ∀I,L1,L2,V. lcoeq 0 0 L1 L2 → lcoeq 0 0 (L1.ⓑ{I}V) (L2.ⓑ{I}V) | lcoeq_pair: ∀I1,I2,L1,L2,V1,V2,e. lcoeq 0 e L1 L2 → lcoeq 0 (⫯e) (L1.ⓑ{I1}V1) (L2.ⓑ{I2}V2) | lcoeq_succ: ∀I,L1,L2,V,d,e. lcoeq d e L1 L2 → lcoeq (⫯d) e (L1.ⓑ{I}V) (L2.ⓑ{I}V) . interpretation "coequivalence (local environment)" 'CoEq d e L1 L2 = (lcoeq d e L1 L2). (* Basic properties *********************************************************) lemma lcoeq_pair_lt: ∀I1,I2,L1,L2,V1,V2,e. L1 ≅[0, ⫰e] L2 → 0 < e → L1.ⓑ{I1}V1 ≅[0, e] L2.ⓑ{I2}V2. #I1 #I2 #L1 #L2 #V1 #V2 #e #HL12 #He <(ylt_inv_O1 … He) /2 width=1 by lcoeq_pair/ qed. lemma lcoeq_succ_lt: ∀I,L1,L2,V,d,e. L1 ≅[⫰d, e] L2 → 0 < d → L1.ⓑ{I}V ≅[d, e] L2. ⓑ{I}V. #I #L1 #L2 #V #d #e #HL12 #Hd <(ylt_inv_O1 … Hd) /2 width=1 by lcoeq_succ/ qed. lemma lcoeq_pair_O_Y: ∀L1,L2. L1 ≅[0, ∞] L2 → ∀I1,I2,V1,V2. L1.ⓑ{I1}V1 ≅[0,∞] L2.ⓑ{I2}V2. #L1 #L2 #HL12 #I1 #I2 #V1 #V2 lapply (lcoeq_pair I1 I2 … V1 V2 … HL12) -HL12 // qed. lemma lcoeq_refl: ∀L,d,e. L ≅[d, e] L. #L elim L -L // #L #I #V #IHL #d elim (ynat_cases … d) [| * #x ] #Hd destruct /2 width=1 by lcoeq_succ/ #e elim (ynat_cases … e) [| * #x ] #He destruct /2 width=1 by lcoeq_zero, lcoeq_pair/ qed. lemma lcoeq_O_Y: ∀L1,L2. |L1| = |L2| → L1 ≅[0, ∞] L2. #L1 elim L1 -L1 [| #L1 #I1 #V1 #IHL1 ] * [2,4: #L2 #I2 #V1 ] normalize /3 width=2 by lcoeq_pair_O_Y/