X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fsyntax%2Ftdeq.ma;h=d431b7677a55bd28dc55f1e47578c4c7d06cbcd9;hp=6e08f744662721eedc90405c94d9abec182ddc24;hb=222044da28742b24584549ba86b1805a87def070;hpb=e2b4ff64df523b4be9d7dc4e92386945846426e7 diff --git a/matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq.ma index 6e08f7446..d431b7677 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq.ma @@ -12,9 +12,9 @@ (* *) (**************************************************************************) -include "basic_2/notation/relations/lazyeq_4.ma". +include "basic_2/notation/relations/stareq_4.ma". include "basic_2/syntax/item_sd.ma". -include "basic_2/syntax/lenv.ma". +include "basic_2/syntax/term.ma". (* DEGREE-BASED EQUIVALENCE ON TERMS ****************************************) @@ -26,15 +26,12 @@ inductive tdeq (h) (o): relation term ≝ . interpretation - "degree-based equivalence (terms)" - 'LazyEq h o T1 T2 = (tdeq h o T1 T2). - -definition cdeq: ∀h. sd h → relation3 lenv term term ≝ - λh,o,L. tdeq h o. + "context-free degree-based equivalence (term)" + 'StarEq h o T1 T2 = (tdeq h o T1 T2). (* Basic inversion lemmas ***************************************************) -fact tdeq_inv_sort1_aux: ∀h,o,X,Y. X ≡[h, o] Y → ∀s1. X = ⋆s1 → +fact tdeq_inv_sort1_aux: ∀h,o,X,Y. X ≛[h, o] Y → ∀s1. X = ⋆s1 → ∃∃s2,d. deg h o s1 d & deg h o s2 d & Y = ⋆s2. #h #o #X #Y * -X -Y [ #s1 #s2 #d #Hs1 #Hs2 #s #H destruct /2 width=5 by ex3_2_intro/ @@ -44,32 +41,32 @@ fact tdeq_inv_sort1_aux: ∀h,o,X,Y. X ≡[h, o] Y → ∀s1. X = ⋆s1 → ] qed-. -lemma tdeq_inv_sort1: ∀h,o,Y,s1. ⋆s1 ≡[h, o] Y → +lemma tdeq_inv_sort1: ∀h,o,Y,s1. ⋆s1 ≛[h, o] Y → ∃∃s2,d. deg h o s1 d & deg h o s2 d & Y = ⋆s2. /2 width=3 by tdeq_inv_sort1_aux/ qed-. -fact tdeq_inv_lref1_aux: ∀h,o,X,Y. X ≡[h, o] Y → ∀i. X = #i → Y = #i. +fact tdeq_inv_lref1_aux: ∀h,o,X,Y. X ≛[h, o] Y → ∀i. X = #i → Y = #i. #h #o #X #Y * -X -Y // [ #s1 #s2 #d #_ #_ #j #H destruct | #I #V1 #V2 #T1 #T2 #_ #_ #j #H destruct ] qed-. -lemma tdeq_inv_lref1: ∀h,o,Y,i. #i ≡[h, o] Y → Y = #i. +lemma tdeq_inv_lref1: ∀h,o,Y,i. #i ≛[h, o] Y → Y = #i. /2 width=5 by tdeq_inv_lref1_aux/ qed-. -fact tdeq_inv_gref1_aux: ∀h,o,X,Y. X ≡[h, o] Y → ∀l. X = §l → Y = §l. +fact tdeq_inv_gref1_aux: ∀h,o,X,Y. X ≛[h, o] Y → ∀l. X = §l → Y = §l. #h #o #X #Y * -X -Y // [ #s1 #s2 #d #_ #_ #k #H destruct | #I #V1 #V2 #T1 #T2 #_ #_ #k #H destruct ] qed-. -lemma tdeq_inv_gref1: ∀h,o,Y,l. §l ≡[h, o] Y → Y = §l. +lemma tdeq_inv_gref1: ∀h,o,Y,l. §l ≛[h, o] Y → Y = §l. /2 width=5 by tdeq_inv_gref1_aux/ qed-. -fact tdeq_inv_pair1_aux: ∀h,o,X,Y. X ≡[h, o] Y → ∀I,V1,T1. X = ②{I}V1.T1 → - ∃∃V2,T2. V1 ≡[h, o] V2 & T1 ≡[h, o] T2 & Y = ②{I}V2.T2. +fact tdeq_inv_pair1_aux: ∀h,o,X,Y. X ≛[h, o] Y → ∀I,V1,T1. X = ②{I}V1.T1 → + ∃∃V2,T2. V1 ≛[h, o] V2 & T1 ≛[h, o] T2 & Y = ②{I}V2.T2. #h #o #X #Y * -X -Y [ #s1 #s2 #d #_ #_ #J #W1 #U1 #H destruct | #i #J #W1 #U1 #H destruct @@ -78,19 +75,19 @@ fact tdeq_inv_pair1_aux: ∀h,o,X,Y. X ≡[h, o] Y → ∀I,V1,T1. X = ②{I}V1. ] qed-. -lemma tdeq_inv_pair1: ∀h,o,I,V1,T1,Y. ②{I}V1.T1 ≡[h, o] Y → - ∃∃V2,T2. V1 ≡[h, o] V2 & T1 ≡[h, o] T2 & Y = ②{I}V2.T2. +lemma tdeq_inv_pair1: ∀h,o,I,V1,T1,Y. ②{I}V1.T1 ≛[h, o] Y → + ∃∃V2,T2. V1 ≛[h, o] V2 & T1 ≛[h, o] T2 & Y = ②{I}V2.T2. /2 width=3 by tdeq_inv_pair1_aux/ qed-. (* Advanced inversion lemmas ************************************************) -lemma tdeq_inv_sort1_deg: ∀h,o,Y,s1. ⋆s1 ≡[h, o] Y → ∀d. deg h o s1 d → +lemma tdeq_inv_sort1_deg: ∀h,o,Y,s1. ⋆s1 ≛[h, o] Y → ∀d. deg h o s1 d → ∃∃s2. deg h o s2 d & Y = ⋆s2. #h #o #Y #s1 #H #d #Hs1 elim (tdeq_inv_sort1 … H) -H #s2 #x #Hx <(deg_mono h o … Hx … Hs1) -s1 -d /2 width=3 by ex2_intro/ qed-. -lemma tdeq_inv_sort_deg: ∀h,o,s1,s2. ⋆s1 ≡[h, o] ⋆s2 → +lemma tdeq_inv_sort_deg: ∀h,o,s1,s2. ⋆s1 ≛[h, o] ⋆s2 → ∀d1,d2. deg h o s1 d1 → deg h o s2 d2 → d1 = d2. #h #o #s1 #y #H #d1 #d2 #Hs1 #Hy @@ -98,13 +95,20 @@ elim (tdeq_inv_sort1_deg … H … Hs1) -s1 #s2 #Hs2 #H destruct <(deg_mono h o … Hy … Hs2) -s2 -d1 // qed-. -lemma tdeq_inv_pair: ∀h,o,I1,I2,V1,V2,T1,T2. ②{I1}V1.T1 ≡[h, o] ②{I2}V2.T2 → - ∧∧ I1 = I2 & V1 ≡[h, o] V2 & T1 ≡[h, o] T2. +lemma tdeq_inv_pair: ∀h,o,I1,I2,V1,V2,T1,T2. ②{I1}V1.T1 ≛[h, o] ②{I2}V2.T2 → + ∧∧ I1 = I2 & V1 ≛[h, o] V2 & T1 ≛[h, o] T2. #h #o #I1 #I2 #V1 #V2 #T1 #T2 #H elim (tdeq_inv_pair1 … H) -H #V0 #T0 #HV #HT #H destruct /2 width=1 by and3_intro/ qed-. -lemma tdeq_inv_pair_xy_y: ∀h,o,I,T,V. ②{I}V.T ≡[h, o] T → ⊥. +lemma tdeq_inv_pair_xy_x: ∀h,o,I,V,T. ②{I}V.T ≛[h, o] V → ⊥. +#h #o #I #V elim V -V +[ #J #T #H elim (tdeq_inv_pair1 … H) -H #X #Y #_ #_ #H destruct +| #J #X #Y #IHX #_ #T #H elim (tdeq_inv_pair … H) -H #H #HY #_ destruct /2 width=2 by/ +] +qed-. + +lemma tdeq_inv_pair_xy_y: ∀h,o,I,T,V. ②{I}V.T ≛[h, o] T → ⊥. #h #o #I #T elim T -T [ #J #V #H elim (tdeq_inv_pair1 … H) -H #X #Y #_ #_ #H destruct | #J #X #Y #_ #IHY #V #H elim (tdeq_inv_pair … H) -H #H #_ #HY destruct /2 width=2 by/ @@ -113,7 +117,7 @@ qed-. (* Basic forward lemmas *****************************************************) -lemma tdeq_fwd_atom1: ∀h,o,I,Y. ⓪{I} ≡[h, o] Y → ∃J. Y = ⓪{J}. +lemma tdeq_fwd_atom1: ∀h,o,I,Y. ⓪{I} ≛[h, o] Y → ∃J. Y = ⓪{J}. #h #o * #x #Y #H [ elim (tdeq_inv_sort1 … H) -H ] /3 width=4 by tdeq_inv_gref1, tdeq_inv_lref1, ex_intro/ qed-. @@ -131,7 +135,7 @@ lemma tdeq_sym: ∀h,o. symmetric … (tdeq h o). /2 width=3 by tdeq_sort, tdeq_lref, tdeq_gref, tdeq_pair/ qed-. -lemma tdeq_dec: ∀h,o,T1,T2. Decidable (tdeq h o T1 T2). +lemma tdeq_dec: ∀h,o,T1,T2. Decidable (T1 ≛[h, o] T2). #h #o #T1 elim T1 -T1 [ * #s1 | #I1 #V1 #T1 #IHV #IHT ] * [1,3,5,7: * #s2 |*: #I2 #V2 #T2 ] [ elim (deg_total h o s1) #d1 #H1 elim (deg_total h o s2) #d2 #H2 @@ -172,10 +176,10 @@ qed-. (* Negated inversion lemmas *************************************************) lemma tdneq_inv_pair: ∀h,o,I1,I2,V1,V2,T1,T2. - (②{I1}V1.T1 ≡[h, o] ②{I2}V2.T2 → ⊥) → + (②{I1}V1.T1 ≛[h, o] ②{I2}V2.T2 → ⊥) → ∨∨ I1 = I2 → ⊥ - | V1 ≡[h, o] V2 → ⊥ - | (T1 ≡[h, o] T2 → ⊥). + | (V1 ≛[h, o] V2 → ⊥) + | (T1 ≛[h, o] T2 → ⊥). #h #o #I1 #I2 #V1 #V2 #T1 #T2 #H12 elim (eq_item2_dec I1 I2) /3 width=1 by or3_intro0/ #H destruct elim (tdeq_dec h o V1 V2) /3 width=1 by or3_intro1/