X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fstatic%2Frdeq.ma;h=2e155d1ee3680677e3591429dba64fe49610c166;hb=a454837a256907d2f83d42ced7be847e10361ea9;hp=6cb56dcaf43200dc2c6eebb6751107f973902db0;hpb=4173283e148199871d787c53c0301891deb90713;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/static_2/static/rdeq.ma b/matita/matita/contribs/lambdadelta/static_2/static/rdeq.ma index 6cb56dcaf..2e155d1ee 100644 --- a/matita/matita/contribs/lambdadelta/static_2/static/rdeq.ma +++ b/matita/matita/contribs/lambdadelta/static_2/static/rdeq.ma @@ -31,8 +31,8 @@ interpretation (* Basic properties ***********************************************************) -lemma frees_tdeq_conf_rdeq: ∀f,L1,T1. L1 ⊢ 𝐅*⦃T1⦄ ≘ f → ∀T2. T1 ≛ T2 → - ∀L2. L1 ≛[f] L2 → L2 ⊢ 𝐅*⦃T2⦄ ≘ f. +lemma frees_tdeq_conf_rdeq: ∀f,L1,T1. L1 ⊢ 𝐅+⦃T1⦄ ≘ f → ∀T2. T1 ≛ T2 → + ∀L2. L1 ≛[f] L2 → L2 ⊢ 𝐅+⦃T2⦄ ≘ f. #f #L1 #T1 #H elim H -f -L1 -T1 [ #f #L1 #s1 #Hf #X #H1 #L2 #_ elim (tdeq_inv_sort1 … H1) -H1 #s2 #H destruct @@ -65,12 +65,12 @@ lemma frees_tdeq_conf_rdeq: ∀f,L1,T1. L1 ⊢ 𝐅*⦃T1⦄ ≘ f → ∀T2. T1 ] qed-. -lemma frees_tdeq_conf: ∀f,L,T1. L ⊢ 𝐅*⦃T1⦄ ≘ f → - ∀T2. T1 ≛ T2 → L ⊢ 𝐅*⦃T2⦄ ≘ f. +lemma frees_tdeq_conf: ∀f,L,T1. L ⊢ 𝐅+⦃T1⦄ ≘ f → + ∀T2. T1 ≛ T2 → L ⊢ 𝐅+⦃T2⦄ ≘ f. /4 width=7 by frees_tdeq_conf_rdeq, sex_refl, ext2_refl/ qed-. -lemma frees_rdeq_conf: ∀f,L1,T. L1 ⊢ 𝐅*⦃T⦄ ≘ f → - ∀L2. L1 ≛[f] L2 → L2 ⊢ 𝐅*⦃T⦄ ≘ f. +lemma frees_rdeq_conf: ∀f,L1,T. L1 ⊢ 𝐅+⦃T⦄ ≘ f → + ∀L2. L1 ≛[f] L2 → L2 ⊢ 𝐅+⦃T⦄ ≘ f. /2 width=7 by frees_tdeq_conf_rdeq, tdeq_refl/ qed-. lemma tdeq_rex_conf (R): s_r_confluent1 … cdeq (rex R). @@ -79,7 +79,7 @@ lemma tdeq_rex_conf (R): s_r_confluent1 … cdeq (rex R). qed-. lemma tdeq_rex_div (R): ∀T1,T2. T1 ≛ T2 → - ∀L1,L2. L1 ⪤[R, T2] L2 → L1 ⪤[R, T1] L2. + ∀L1,L2. L1 ⪤[R,T2] L2 → L1 ⪤[R,T1] L2. /3 width=5 by tdeq_rex_conf, tdeq_sym/ qed-. lemma tdeq_rdeq_conf: s_r_confluent1 … cdeq rdeq. @@ -100,7 +100,7 @@ lemma rdeq_pair: ∀I,L1,L2,V1,V2. L1 ≛[V1] L2 → V1 ≛ V2 → L1.ⓑ{I}V1 ≛[#0] L2.ⓑ{I}V2. /2 width=1 by rex_pair/ qed. (* -lemma rdeq_unit: ∀f,I,L1,L2. 𝐈⦃f⦄ → L1 ⪤[cdeq_ext, cfull, f] L2 → +lemma rdeq_unit: ∀f,I,L1,L2. 𝐈⦃f⦄ → L1 ⪤[cdeq_ext,cfull,f] L2 → L1.ⓤ{I} ≛[#0] L2.ⓤ{I}. /2 width=3 by rex_unit/ qed. *) @@ -130,7 +130,7 @@ lemma rdeq_inv_zero: ∀Y1,Y2. Y1 ≛[#0] Y2 → ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆ | ∃∃I,L1,L2,V1,V2. L1 ≛[V1] L2 & V1 ≛ V2 & Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2 - | ∃∃f,I,L1,L2. 𝐈⦃f⦄ & L1 ⪤[cdeq_ext h o, cfull, f] L2 & + | ∃∃f,I,L1,L2. 𝐈⦃f⦄ & L1 ⪤[cdeq_ext h o,cfull,f] L2 & Y1 = L1.ⓤ{I} & Y2 = L2.ⓤ{I}. #Y1 #Y2 #H elim (rex_inv_zero … H) -H * /3 width=9 by or3_intro0, or3_intro1, or3_intro2, ex4_5_intro, ex4_4_intro, conj/