]> matita.cs.unibo.it Git - helm.git/commitdiff
diamond property of reduction!
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Thu, 15 Sep 2016 15:28:39 +0000 (15:28 +0000)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Thu, 15 Sep 2016 15:28:39 +0000 (15:28 +0000)
13 files changed:
matita/matita/contribs/lambdadelta/basic_2/relocation/lexs.ma
matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpr.ma
matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpr.ma
matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpr_drops.ma
matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpr_lfpr.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx.ma
matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_drops.ma
matita/matita/contribs/lambdadelta/basic_2/rt_transition/lpr_lpr.ma [deleted file]
matita/matita/contribs/lambdadelta/basic_2/rt_transition/partial.txt
matita/matita/contribs/lambdadelta/basic_2/static/lfxs.ma
matita/matita/contribs/lambdadelta/basic_2/static/lfxs_drops.ma
matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lfxs.ma
matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl

index 7fd35071acaea51d259b3f7f828de5ab4ee4970a..3a779a649182b05ca4917e1aa643e21f6ec2c6c4 100644 (file)
@@ -181,6 +181,22 @@ lemma lexs_refl: ∀RN,RP,f.
 #L #I #V #IH * * /2 width=1 by lexs_next, lexs_push/
 qed.
 
+lemma lexs_pair_repl: ∀RN,RP,f,I,L1,L2,V1,V2.
+                      L1.ⓑ{I}V1 ⦻*[RN, RP, f] L2.ⓑ{I}V2 →
+                      ∀W1,W2. RN L1 W1 W2 → RP L1 W1 W2 →
+                      L1.ⓑ{I}W1 ⦻*[RN, RP, f] L2.ⓑ{I}W2.
+#RN #RP #f #I #L1 #L2 #V1 #V2 #HL12 #W1 #W2 #HN #HP
+elim (lexs_fwd_pair … HL12) -HL12 /2 width=1 by lexs_inv_tl/
+qed-.
+
+lemma lexs_co: ∀RN1,RP1,RN2,RP2.
+               (∀L1,T1,T2. RN1 L1 T1 T2 → RN2 L1 T1 T2) →
+               (∀L1,T1,T2. RP1 L1 T1 T2 → RP2 L1 T1 T2) →
+               ∀f,L1,L2. L1 ⦻*[RN1, RP1, f] L2 → L1 ⦻*[RN2, RP2, f] L2.
+#RN1 #RP1 #RN2 #RP2 #HRN #HRP #f #L1 #L2 #H elim H -f -L1 -L2
+/3 width=1 by lexs_atom, lexs_next, lexs_push/
+qed-.
+
 lemma sle_lexs_trans: ∀RN,RP. (∀L,T1,T2. RN L T1 T2 → RP L T1 T2) →
                       ∀f2,L1,L2. L1 ⦻*[RN, RP, f2] L2 →
                       ∀f1. f1 ⊆ f2 → L1 ⦻*[RN, RP, f1] L2.
@@ -206,11 +222,3 @@ lemma sle_lexs_conf: ∀RN,RP. (∀L,T1,T2. RP L T1 T2 → RN L T1 T2) →
   #g2 #H #H2 destruct /3 width=5 by lexs_next/
 ]
 qed-.
-
-lemma lexs_co: ∀RN1,RP1,RN2,RP2.
-               (∀L1,T1,T2. RN1 L1 T1 T2 → RN2 L1 T1 T2) →
-               (∀L1,T1,T2. RP1 L1 T1 T2 → RP2 L1 T1 T2) →
-               ∀f,L1,L2. L1 ⦻*[RN1, RP1, f] L2 → L1 ⦻*[RN2, RP2, f] L2.
-#RN1 #RP1 #RN2 #RP2 #HRN #HRP #f #L1 #L2 #H elim H -f -L1 -L2
-/3 width=1 by lexs_atom, lexs_next, lexs_push/
-qed-.
index a5d3ca4d3188f38ef9019301d9dfe924108d4c35..6bcfa173e8ec96dfc351c475a5820c71c14026ed 100644 (file)
@@ -107,7 +107,3 @@ qed-.
             pr2_gen_csort pr2_gen_cflat pr2_gen_cbind
             pr2_gen_ctail pr2_ctail
 *)
-(* Basic_1: removed local theorems 4:
-            pr0_delta_eps pr0_cong_delta
-            pr2_free_free pr2_free_delta
-*)
index a38aff49da8083feb56e641e9e95e34fcfd0db98..299b0f8c8139115f3b49db70aac74330421fddc2 100644 (file)
@@ -46,6 +46,12 @@ lemma lfpr_gref: ∀h,I,G,L1,L2,V1,V2,l.
                  ⦃G, L1⦄ ⊢ ➡[h, §l] L2 → ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡[h, §l] L2.ⓑ{I}V2.
 /2 width=1 by lfxs_gref/ qed.
 
+lemma lfpr_pair_repl_dx: ∀h,I,G,L1,L2,T,V,V1.
+                         ⦃G, L1.ⓑ{I}V⦄ ⊢ ➡[h, T] L2.ⓑ{I}V1 →
+                         ∀V2. ⦃G, L1⦄ ⊢ V ➡[h] V2 →
+                         ⦃G, L1.ⓑ{I}V⦄ ⊢ ➡[h, T] L2.ⓑ{I}V2.
+/2 width=2 by lfxs_pair_repl_dx/ qed-.
+
 (* Basic inversion lemmas ***************************************************)
 
 lemma lfpr_inv_atom_sn: ∀h,I,G,Y2. ⦃G, ⋆⦄ ⊢ ➡[h, ⓪{I}] Y2 → Y2 = ⋆.
@@ -117,14 +123,16 @@ lemma lfpr_fwd_pair_sn: ∀h,I,G,L1,L2,V,T.
                         ⦃G, L1⦄ ⊢ ➡[h, ②{I}V.T] L2 → ⦃G, L1⦄ ⊢ ➡[h, V] L2.
 /2 width=3 by lfxs_fwd_pair_sn/ qed-.
 
-(* Basic_2A1: removed theorems 11:
+(* Basic_2A1: removed theorems 14:
               lpr_inv_atom1 lpr_inv_pair1 lpr_inv_atom2 lpr_inv_pair2
               lpr_refl lpr_pair
               lpr_fwd_length lpr_lpx
               lpr_drop_conf drop_lpr_trans lpr_drop_trans_O1
+              cpr_conf_lpr lpr_cpr_conf_dx lpr_cpr_conf_sn
 *)
-(* Basic_1: removed theorems 7: wcpr0_gen_sort wcpr0_gen_head
-                                wcpr0_getl wcpr0_getl_back
-                                pr0_subst1_back
-                                wcpr0_drop wcpr0_drop_back
+(* Basic_1: removed theorems 7:
+            wcpr0_gen_sort wcpr0_gen_head
+            wcpr0_getl wcpr0_getl_back
+            pr0_subst1_back
+            wcpr0_drop wcpr0_drop_back
 *)
index b3deebd859e0ea411d626ca095f62897ec2f578d..3156e7252cf303c9721374975a84951a3919e93f 100644 (file)
@@ -30,3 +30,11 @@ lemma lfpr_drops_conf: ∀h,G. dropable_sn (cpm 0 h G).
 
 lemma lfpr_drops_trans: ∀h,G. dropable_dx (cpm 0 h G).
 /2 width=5 by lfxs_dropable_dx/ qed-.
+
+lemma lfpr_inv_lref_sn: ∀h,G,L1,L2,i. ⦃G, L1⦄ ⊢ ➡[h, #i] L2 → ∀I,K1,V1. ⬇*[i] L1 ≡ K1.ⓑ{I}V1 →
+                        ∃∃K2,V2. ⬇*[i] L2 ≡ K2.ⓑ{I}V2 & ⦃G, K1⦄ ⊢ ➡[h, V1] K2 & ⦃G, K1⦄ ⊢ V1 ➡[h] V2.
+/2 width=3 by lfxs_inv_lref_sn/ qed-.
+
+lemma lfpr_inv_lref_dx: ∀h,G,L1,L2,i. ⦃G, L1⦄ ⊢ ➡[h, #i] L2 → ∀I,K2,V2. ⬇*[i] L2 ≡ K2.ⓑ{I}V2 →
+                        ∃∃K1,V1. ⬇*[i] L1 ≡ K1.ⓑ{I}V1 & ⦃G, K1⦄ ⊢ ➡[h, V1] K2 & ⦃G, K1⦄ ⊢ V1 ➡[h] V2.
+/2 width=3 by lfxs_inv_lref_dx/ qed-.
diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpr_lfpr.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpr_lfpr.ma
new file mode 100644 (file)
index 0000000..e8955fd
--- /dev/null
@@ -0,0 +1,383 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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 "basic_2/rt_transition/cpm_lsubr.ma".
+include "basic_2/rt_transition/cpr.ma".
+include "basic_2/rt_transition/cpr_drops.ma".
+include "basic_2/rt_transition/lfpr_drops.ma".
+include "basic_2/rt_transition/lfpr_fqup.ma".
+
+(* PARALLEL R-TRANSITION FOR LOCAL ENV.S ON REFERRED ENTRIES ****************)
+
+(* Main properties with context-sensitive parallel r-transition for terms ***)
+
+fact cpr_conf_lfpr_atom_atom:
+   ∀h,I,G,L1,L2. ∃∃T. ⦃G, L1⦄ ⊢ ⓪{I} ➡[h] T & ⦃G, L2⦄ ⊢ ⓪{I} ➡[h] T.
+/2 width=3 by ex2_intro/ qed-.
+
+fact cpr_conf_lfpr_atom_delta:
+   ∀h,G,L0,i. (
+      ∀L,T. ⦃G, L0, #i⦄ ⊐+ ⦃G, L, T⦄ →
+      ∀T1. ⦃G, L⦄ ⊢ T ➡[h] T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡[h] T2 →
+      ∀L1. ⦃G, L⦄ ⊢ ➡[h, T] L1 → ∀L2. ⦃G, L⦄ ⊢ ➡[h, T] L2 →
+      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡[h] T0 & ⦃G, L2⦄ ⊢ T2 ➡[h] T0
+   ) →
+   ∀K0,V0. ⬇*[i] L0 ≡ K0.ⓓV0 →
+   ∀V2. ⦃G, K0⦄ ⊢ V0 ➡[h] V2 → ∀T2. ⬆*[⫯i] V2 ≡ T2 →
+   ∀L1. ⦃G, L0⦄ ⊢ ➡[h, #i] L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡[h, #i] L2 →
+   ∃∃T. ⦃G, L1⦄ ⊢ #i ➡[h] T & ⦃G, L2⦄ ⊢ T2 ➡[h] T.
+#h #G #L0 #i #IH #K0 #V0 #HLK0 #V2 #HV02 #T2 #HVT2 #L1 #HL01 #L2 #HL02
+elim (lfpr_inv_lref_sn … HL01 … HLK0) -HL01 #K1 #V1 #HLK1 #HK01 #HV01
+elim (lfpr_inv_lref_sn … HL02 … HLK0) -HL02 #K2 #W2 #HLK2 #HK02 #_
+lapply (drops_isuni_fwd_drop2 … HLK2) // -W2 #HLK2
+lapply (fqup_lref … G … HLK0) -HLK0 #HLK0
+elim (IH … HLK0 … HV01 … HV02 … HK01 … HK02) -L0 -K0 -V0 #V #HV1 #HV2
+elim (cpm_lifts … HV2 … HLK2 … HVT2) -K2 -V2
+/3 width=6 by cpm_delta_drops, ex2_intro/
+qed-.
+
+fact cpr_conf_lfpr_delta_delta:
+   ∀h,G,L0,i. (
+      ∀L,T. ⦃G, L0, #i⦄ ⊐+ ⦃G, L, T⦄ →
+      ∀T1. ⦃G, L⦄ ⊢ T ➡[h] T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡[h] T2 →
+      ∀L1. ⦃G, L⦄ ⊢ ➡[h, T] L1 → ∀L2. ⦃G, L⦄ ⊢ ➡[h, T] L2 →
+      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡[h] T0 & ⦃G, L2⦄ ⊢ T2 ➡[h] T0
+   ) →
+   ∀K0,V0. ⬇*[i] L0 ≡ K0.ⓓV0 →
+   ∀V1. ⦃G, K0⦄ ⊢ V0 ➡[h] V1 → ∀T1. ⬆*[⫯i] V1 ≡ T1 →
+   ∀KX,VX. ⬇*[i] L0 ≡ KX.ⓓVX →
+   ∀V2. ⦃G, KX⦄ ⊢ VX ➡[h] V2 → ∀T2. ⬆*[⫯i] V2 ≡ T2 →
+   ∀L1. ⦃G, L0⦄ ⊢ ➡[h, #i] L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡[h, #i] L2 →
+   ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡[h] T & ⦃G, L2⦄ ⊢ T2 ➡[h] T.
+#h #G #L0 #i #IH #K0 #V0 #HLK0 #V1 #HV01 #T1 #HVT1
+#KX #VX #H #V2 #HV02 #T2 #HVT2 #L1 #HL01 #L2 #HL02
+lapply (drops_mono … H … HLK0) -H #H destruct
+elim (lfpr_inv_lref_sn … HL01 … HLK0) -HL01 #K1 #W1 #HLK1 #HK01 #_
+lapply (drops_isuni_fwd_drop2 … HLK1) -W1 // #HLK1
+elim (lfpr_inv_lref_sn … HL02 … HLK0) -HL02 #K2 #W2 #HLK2 #HK02 #_
+lapply (drops_isuni_fwd_drop2 … HLK2) -W2 // #HLK2
+lapply (fqup_lref … G … HLK0) -HLK0 #HLK0
+elim (IH … HLK0 … HV01 … HV02 … HK01 … HK02) -L0 -K0 -V0 #V #HV1 #HV2
+elim (cpm_lifts … HV1 … HLK1 … HVT1) -K1 -V1 #T #HVT #HT1
+elim (cpm_lifts … HV2 … HLK2 … HVT2) -K2 -V2 #X #HX #HT2
+lapply (lifts_mono … HX … HVT) #H destruct
+/2 width=3 by ex2_intro/
+qed-.
+
+fact cpr_conf_lfpr_bind_bind:
+   ∀h,p,I,G,L0,V0,T0. (
+      ∀L,T. ⦃G, L0, ⓑ{p,I}V0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+      ∀T1. ⦃G, L⦄ ⊢ T ➡[h] T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡[h] T2 →
+      ∀L1. ⦃G, L⦄ ⊢ ➡[h, T] L1 → ∀L2. ⦃G, L⦄ ⊢ ➡[h, T] L2 →
+      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡[h] T0 & ⦃G, L2⦄ ⊢ T2 ➡[h] T0
+   ) →
+   ∀V1. ⦃G, L0⦄ ⊢ V0 ➡[h] V1 → ∀T1. ⦃G, L0.ⓑ{I}V0⦄ ⊢ T0 ➡[h] T1 →
+   ∀V2. ⦃G, L0⦄ ⊢ V0 ➡[h] V2 → ∀T2. ⦃G, L0.ⓑ{I}V0⦄ ⊢ T0 ➡[h] T2 →
+   ∀L1. ⦃G, L0⦄ ⊢ ➡[h, ⓑ{p,I}V0.T0] L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡[h, ⓑ{p,I}V0.T0] L2 →
+   ∃∃T. ⦃G, L1⦄ ⊢ ⓑ{p,I}V1.T1 ➡[h] T & ⦃G, L2⦄ ⊢ ⓑ{p,I}V2.T2 ➡[h] T.
+#h #p #I #G #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01
+#V2 #HV02 #T2 #HT02 #L1 #HL01 #L2 #HL02
+elim (lfpr_inv_bind … HL01) -HL01 #H1V0 #H1T0
+elim (lfpr_inv_bind … HL02) -HL02 #H2V0 #H2T0
+elim (IH … HV01 … HV02 … H1V0 … H2V0) //
+elim (IH … HT01 … HT02 (L1.ⓑ{I}V1) … (L2.ⓑ{I}V2)) -IH
+/3 width=5 by lfpr_pair_repl_dx, cpm_bind, ex2_intro/
+qed-.
+
+fact cpr_conf_lfpr_bind_zeta:
+   ∀h,G,L0,V0,T0. (
+      ∀L,T. ⦃G, L0, +ⓓV0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+      ∀T1. ⦃G, L⦄ ⊢ T ➡[h] T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡[h] T2 →
+      ∀L1. ⦃G, L⦄ ⊢ ➡[h, T] L1 → ∀L2. ⦃G, L⦄ ⊢ ➡[h, T] L2 →
+      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡[h] T0 & ⦃G, L2⦄ ⊢ T2 ➡[h] T0
+   ) →
+   ∀V1. ⦃G, L0⦄ ⊢ V0 ➡[h] V1 → ∀T1. ⦃G, L0.ⓓV0⦄ ⊢ T0 ➡[h] T1 →
+   ∀T2. ⦃G, L0.ⓓV0⦄ ⊢ T0 ➡[h] T2 → ∀X2. ⬆*[1] X2 ≡ T2 →
+   ∀L1. ⦃G, L0⦄ ⊢ ➡[h, +ⓓV0.T0] L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡[h, +ⓓV0.T0] L2 →
+   ∃∃T. ⦃G, L1⦄ ⊢ +ⓓV1.T1 ➡[h] T & ⦃G, L2⦄ ⊢ X2 ➡[h] T.
+#h #G #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01
+#T2 #HT02 #X2 #HXT2 #L1 #HL01 #L2 #HL02
+elim (lfpr_inv_bind … HL01) -HL01 #H1V0 #H1T0
+elim (lfpr_inv_bind … HL02) -HL02 #H2V0 #H2T0
+elim (IH … HT01 … HT02 (L1.ⓓV1) … (L2.ⓓV1)) -IH -HT01 -HT02 /2 width=4 by lfpr_pair_repl_dx/ -L0 -V0 -T0 #T #HT1 #HT2
+elim (cpm_inv_lifts1 … HT2 … L2 … HXT2) -T2 /3 width=3 by drops_refl, drops_drop, cpm_zeta, ex2_intro/
+qed-.
+
+fact cpr_conf_lfpr_zeta_zeta:
+   ∀h,G,L0,V0,T0. (
+      ∀L,T. ⦃G, L0, +ⓓV0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+      ∀T1. ⦃G, L⦄ ⊢ T ➡[h] T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡[h] T2 →
+      ∀L1. ⦃G, L⦄ ⊢ ➡[h, T] L1 → ∀L2. ⦃G, L⦄ ⊢ ➡[h, T] L2 →
+      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡[h] T0 & ⦃G, L2⦄ ⊢ T2 ➡[h] T0
+   ) →
+   ∀T1. ⦃G, L0.ⓓV0⦄ ⊢ T0 ➡[h] T1 → ∀X1. ⬆*[1] X1 ≡ T1 →
+   ∀T2. ⦃G, L0.ⓓV0⦄ ⊢ T0 ➡[h] T2 → ∀X2. ⬆*[1] X2 ≡ T2 →
+   ∀L1. ⦃G, L0⦄ ⊢ ➡[h, +ⓓV0.T0] L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡[h, +ⓓV0.T0] L2 →
+   ∃∃T. ⦃G, L1⦄ ⊢ X1 ➡[h] T & ⦃G, L2⦄ ⊢ X2 ➡[h] T.
+#h #G #L0 #V0 #T0 #IH #T1 #HT01 #X1 #HXT1
+#T2 #HT02 #X2 #HXT2 #L1 #HL01 #L2 #HL02
+elim (lfpr_inv_bind … HL01) -HL01 #H1V0 #H1T0
+elim (lfpr_inv_bind … HL02) -HL02 #H2V0 #H2T0
+elim (IH … HT01 … HT02 (L1.ⓓV0) … (L2.ⓓV0)) -IH -HT01 -HT02 /2 width=4 by lfpr_pair_repl_dx/ -L0 -T0 #T #HT1 #HT2
+elim (cpm_inv_lifts1 … HT1 … L1 … HXT1) -T1 /3 width=2 by drops_refl, drops_drop/ #T1 #HT1 #HXT1
+elim (cpm_inv_lifts1 … HT2 … L2 … HXT2) -T2 /3 width=2 by drops_refl, drops_drop/ #T2 #HT2 #HXT2
+lapply (lifts_inj … HT2 … HT1) -T #H destruct /2 width=3 by ex2_intro/
+qed-.
+
+fact cpr_conf_lfpr_flat_flat:
+   ∀h,I,G,L0,V0,T0. (
+      ∀L,T. ⦃G, L0, ⓕ{I}V0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+      ∀T1. ⦃G, L⦄ ⊢ T ➡[h] T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡[h] T2 →
+      ∀L1. ⦃G, L⦄ ⊢ ➡[h, T] L1 → ∀L2. ⦃G, L⦄ ⊢ ➡[h, T] L2 →
+      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡[h] T0 & ⦃G, L2⦄ ⊢ T2 ➡[h] T0
+   ) →
+   ∀V1. ⦃G, L0⦄ ⊢ V0 ➡[h] V1 → ∀T1. ⦃G, L0⦄ ⊢ T0 ➡[h] T1 →
+   ∀V2. ⦃G, L0⦄ ⊢ V0 ➡[h] V2 → ∀T2. ⦃G, L0⦄ ⊢ T0 ➡[h] T2 →
+   ∀L1. ⦃G, L0⦄ ⊢ ➡[h, ⓕ{I}V0.T0] L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡[h, ⓕ{I}V0.T0] L2 →
+   ∃∃T. ⦃G, L1⦄ ⊢ ⓕ{I}V1.T1 ➡[h] T & ⦃G, L2⦄ ⊢ ⓕ{I}V2.T2 ➡[h] T.
+#h #I #G #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01
+#V2 #HV02 #T2 #HT02 #L1 #HL01 #L2 #HL02
+elim (lfpr_inv_flat … HL01) -HL01 #H1V0 #H1T0
+elim (lfpr_inv_flat … HL02) -HL02 #H2V0 #H2T0
+elim (IH … HV01 … HV02 … H1V0 … H2V0) //
+elim (IH … HT01 … HT02 … H1T0 … H2T0) /3 width=5 by cpr_flat, ex2_intro/
+qed-.
+
+fact cpr_conf_lfpr_flat_epsilon:
+   ∀h,G,L0,V0,T0. (
+      ∀L,T. ⦃G, L0, ⓝV0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+      ∀T1. ⦃G, L⦄ ⊢ T ➡[h] T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡[h] T2 →
+      ∀L1. ⦃G, L⦄ ⊢ ➡[h, T] L1 → ∀L2. ⦃G, L⦄ ⊢ ➡[h, T] L2 →
+      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡[h] T0 & ⦃G, L2⦄ ⊢ T2 ➡[h] T0
+   ) →
+   ∀V1,T1. ⦃G, L0⦄ ⊢ T0 ➡[h] T1 → ∀T2. ⦃G, L0⦄ ⊢ T0 ➡[h] T2 →
+   ∀L1. ⦃G, L0⦄ ⊢ ➡[h, ⓝV0.T0] L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡[h, ⓝV0.T0] L2 →
+   ∃∃T. ⦃G, L1⦄ ⊢ ⓝV1.T1 ➡[h] T & ⦃G, L2⦄ ⊢ T2 ➡[h] T.
+#h #G #L0 #V0 #T0 #IH #V1 #T1 #HT01
+#T2 #HT02 #L1 #HL01 #L2 #HL02
+elim (lfpr_inv_flat … HL01) -HL01 #_ #H1T0
+elim (lfpr_inv_flat … HL02) -HL02 #_ #H2T0
+elim (IH … HT01 … HT02 … H1T0 … H2T0) // -L0 -V0 -T0 /3 width=3 by cpm_eps, ex2_intro/
+qed-.
+
+fact cpr_conf_lfpr_epsilon_epsilon:
+   ∀h,G,L0,V0,T0. (
+      ∀L,T. ⦃G, L0, ⓝV0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+      ∀T1. ⦃G, L⦄ ⊢ T ➡[h] T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡[h] T2 →
+      ∀L1. ⦃G, L⦄ ⊢ ➡[h, T] L1 → ∀L2. ⦃G, L⦄ ⊢ ➡[h, T] L2 →
+      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡[h] T0 & ⦃G, L2⦄ ⊢ T2 ➡[h] T0
+   ) →
+   ∀T1. ⦃G, L0⦄ ⊢ T0 ➡[h] T1 → ∀T2. ⦃G, L0⦄ ⊢ T0 ➡[h] T2 →
+   ∀L1. ⦃G, L0⦄ ⊢ ➡[h, ⓝV0.T0] L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡[h, ⓝV0.T0] L2 →
+   ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡[h] T & ⦃G, L2⦄ ⊢ T2 ➡[h] T.
+#h #G #L0 #V0 #T0 #IH #T1 #HT01
+#T2 #HT02 #L1 #HL01 #L2 #HL02
+elim (lfpr_inv_flat … HL01) -HL01 #_ #H1T0
+elim (lfpr_inv_flat … HL02) -HL02 #_ #H2T0
+elim (IH … HT01 … HT02 … H1T0 … H2T0) // -L0 -V0 -T0 /2 width=3 by ex2_intro/
+qed-.
+
+fact cpr_conf_lfpr_flat_beta:
+   ∀h,p,G,L0,V0,W0,T0. (
+      ∀L,T. ⦃G, L0, ⓐV0.ⓛ{p}W0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+      ∀T1. ⦃G, L⦄ ⊢ T ➡[h] T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡[h] T2 →
+      ∀L1. ⦃G, L⦄ ⊢ ➡[h, T] L1 → ∀L2. ⦃G, L⦄ ⊢ ➡[h, T] L2 →
+      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡[h] T0 & ⦃G, L2⦄ ⊢ T2 ➡[h] T0
+   ) →
+   ∀V1. ⦃G, L0⦄ ⊢ V0 ➡[h] V1 → ∀T1. ⦃G, L0⦄ ⊢ ⓛ{p}W0.T0 ➡[h] T1 →
+   ∀V2. ⦃G, L0⦄ ⊢ V0 ➡[h] V2 → ∀W2. ⦃G, L0⦄ ⊢ W0 ➡[h] W2 → ∀T2. ⦃G, L0.ⓛW0⦄ ⊢ T0 ➡[h] T2 →
+   ∀L1. ⦃G, L0⦄ ⊢ ➡[h, ⓐV0.ⓛ{p}W0.T0] L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡[h, ⓐV0.ⓛ{p}W0.T0] L2 →
+   ∃∃T. ⦃G, L1⦄ ⊢ ⓐV1.T1 ➡[h] T & ⦃G, L2⦄ ⊢ ⓓ{p}ⓝW2.V2.T2 ➡[h] T.
+#h #p #G #L0 #V0 #W0 #T0 #IH #V1 #HV01 #X #H
+#V2 #HV02 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
+elim (cpm_inv_abst1 … H) -H #W1 #T1 #HW01 #HT01 #H destruct
+elim (lfpr_inv_flat … HL01) -HL01 #H1V0 #HL01
+elim (lfpr_inv_bind … HL01) -HL01 #H1W0 #H1T0
+elim (lfpr_inv_flat … HL02) -HL02 #H2V0 #HL02
+elim (lfpr_inv_bind … HL02) -HL02 #H2W0 #H2T0
+elim (IH … HV01 … HV02 … H1V0 … H2V0) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
+elim (IH … HW01 … HW02 … H1W0 … H2W0) /2 width=1 by/ #W #HW1 #HW2
+elim (IH … HT01 … HT02 (L1.ⓛW1) … (L2.ⓛW2)) /2 width=4 by lfpr_pair_repl_dx/ -L0 -V0 -W0 -T0 #T #HT1 #HT2
+lapply (lsubr_cpm_trans … HT2 (L2.ⓓⓝW2.V2) ?) -HT2 /2 width=1 by lsubr_beta/ (**) (* full auto not tried *)
+/4 width=5 by cpm_bind, cpr_flat, cpm_beta, ex2_intro/
+qed-.
+
+fact cpr_conf_lfpr_flat_theta:
+   ∀h,p,G,L0,V0,W0,T0. (
+      ∀L,T. ⦃G, L0, ⓐV0.ⓓ{p}W0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+      ∀T1. ⦃G, L⦄ ⊢ T ➡[h] T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡[h] T2 →
+      ∀L1. ⦃G, L⦄ ⊢ ➡[h, T] L1 → ∀L2. ⦃G, L⦄ ⊢ ➡[h, T] L2 →
+      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡[h] T0 & ⦃G, L2⦄ ⊢ T2 ➡[h] T0
+   ) →
+   ∀V1. ⦃G, L0⦄ ⊢ V0 ➡[h] V1 → ∀T1. ⦃G, L0⦄ ⊢ ⓓ{p}W0.T0 ➡[h] T1 →
+   ∀V2. ⦃G, L0⦄ ⊢ V0 ➡[h] V2 → ∀U2. ⬆*[1] V2 ≡ U2 →
+   ∀W2. ⦃G, L0⦄ ⊢ W0 ➡[h] W2 → ∀T2. ⦃G, L0.ⓓW0⦄ ⊢ T0 ➡[h] T2 →
+   ∀L1. ⦃G, L0⦄ ⊢ ➡[h, ⓐV0.ⓓ{p}W0.T0] L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡[h, ⓐV0.ⓓ{p}W0.T0] L2 →
+   ∃∃T. ⦃G, L1⦄ ⊢ ⓐV1.T1 ➡[h] T & ⦃G, L2⦄ ⊢ ⓓ{p}W2.ⓐU2.T2 ➡[h] T.
+#h #p #G #L0 #V0 #W0 #T0 #IH #V1 #HV01 #X #H
+#V2 #HV02 #U2 #HVU2 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
+elim (lfpr_inv_flat … HL01) -HL01 #H1V0 #HL01
+elim (lfpr_inv_bind … HL01) -HL01 #H1W0 #H1T0
+elim (lfpr_inv_flat … HL02) -HL02 #H2V0 #HL02
+elim (lfpr_inv_bind … HL02) -HL02 #H2W0 #H2T0
+elim (IH … HV01 … HV02 … H1V0 … H2V0) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
+elim (cpm_lifts … HV2 … (L2.ⓓW2) … HVU2) -HVU2 /3 width=2 by drops_refl, drops_drop/ #U #HVU #HU2
+elim (cpm_inv_abbr1 … H) -H *
+[ #W1 #T1 #HW01 #HT01 #H destruct
+  elim (IH … HW01 … HW02 … H1W0 … H2W0) /2 width=1 by/
+  elim (IH … HT01 … HT02 (L1.ⓓW1) … (L2.ⓓW2)) /2 width=4 by lfpr_pair_repl_dx/ -L0 -V0 -W0 -T0
+  /4 width=7 by cpm_bind, cpr_flat, cpm_theta, ex2_intro/
+| #T1 #HT01 #HXT1 #H destruct
+  elim (IH … HT01 … HT02 (L1.ⓓW2) … (L2.ⓓW2)) /2 width=4 by lfpr_pair_repl_dx/ -L0 -V0 -W0 -T0 #T #HT1 #HT2
+  elim (cpm_inv_lifts1 … HT1 … L1 … HXT1) -HXT1
+  /4 width=9 by cpr_flat, cpm_zeta, drops_refl, drops_drop, lifts_flat, ex2_intro/
+]
+qed-.
+
+fact cpr_conf_lfpr_beta_beta:
+   ∀h,p,G,L0,V0,W0,T0. (
+      ∀L,T. ⦃G, L0, ⓐV0.ⓛ{p}W0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+      ∀T1. ⦃G, L⦄ ⊢ T ➡[h] T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡[h] T2 →
+      ∀L1. ⦃G, L⦄ ⊢ ➡[h, T] L1 → ∀L2. ⦃G, L⦄ ⊢ ➡[h, T] L2 →
+      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡[h] T0 & ⦃G, L2⦄ ⊢ T2 ➡[h] T0
+   ) →
+   ∀V1. ⦃G, L0⦄ ⊢ V0 ➡[h] V1 → ∀W1. ⦃G, L0⦄ ⊢ W0 ➡[h] W1 → ∀T1. ⦃G, L0.ⓛW0⦄ ⊢ T0 ➡[h] T1 →
+   ∀V2. ⦃G, L0⦄ ⊢ V0 ➡[h] V2 → ∀W2. ⦃G, L0⦄ ⊢ W0 ➡[h] W2 → ∀T2. ⦃G, L0.ⓛW0⦄ ⊢ T0 ➡[h] T2 →
+   ∀L1. ⦃G, L0⦄ ⊢ ➡[h, ⓐV0.ⓛ{p}W0.T0] L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡[h, ⓐV0.ⓛ{p}W0.T0] L2 →
+   ∃∃T. ⦃G, L1⦄ ⊢ ⓓ{p}ⓝW1.V1.T1 ➡[h] T & ⦃G, L2⦄ ⊢ ⓓ{p}ⓝW2.V2.T2 ➡[h] T.
+#h #p #G #L0 #V0 #W0 #T0 #IH #V1 #HV01 #W1 #HW01 #T1 #HT01
+#V2 #HV02 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
+elim (lfpr_inv_flat … HL01) -HL01 #H1V0 #HL01
+elim (lfpr_inv_bind … HL01) -HL01 #H1W0 #H1T0
+elim (lfpr_inv_flat … HL02) -HL02 #H2V0 #HL02
+elim (lfpr_inv_bind … HL02) -HL02 #H2W0 #H2T0
+elim (IH … HV01 … HV02 … H1V0 … H2V0) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
+elim (IH … HW01 … HW02 … H1W0 … H2W0) /2 width=1/ #W #HW1 #HW2
+elim (IH … HT01 … HT02 (L1.ⓛW1) … (L2.ⓛW2)) /2 width=4 by lfpr_pair_repl_dx/ -L0 -V0 -W0 -T0 #T #HT1 #HT2
+lapply (lsubr_cpm_trans … HT1 (L1.ⓓⓝW1.V1) ?) -HT1 /2 width=1 by lsubr_beta/
+lapply (lsubr_cpm_trans … HT2 (L2.ⓓⓝW2.V2) ?) -HT2 /2 width=1 by lsubr_beta/
+/4 width=5 by cpm_bind, cpr_flat, ex2_intro/ (**) (* full auto not tried *)
+qed-.
+
+fact cpr_conf_lfpr_theta_theta:
+   ∀h,p,G,L0,V0,W0,T0. (
+      ∀L,T. ⦃G, L0, ⓐV0.ⓓ{p}W0.T0⦄ ⊐+ ⦃G, L, T⦄ →
+      ∀T1. ⦃G, L⦄ ⊢ T ➡[h] T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡[h] T2 →
+      ∀L1. ⦃G, L⦄ ⊢ ➡[h, T] L1 → ∀L2. ⦃G, L⦄ ⊢ ➡[h, T] L2 →
+      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡[h] T0 & ⦃G, L2⦄ ⊢ T2 ➡[h] T0
+   ) →
+   ∀V1. ⦃G, L0⦄ ⊢ V0 ➡[h] V1 → ∀U1. ⬆*[1] V1 ≡ U1 →
+   ∀W1. ⦃G, L0⦄ ⊢ W0 ➡[h] W1 → ∀T1. ⦃G, L0.ⓓW0⦄ ⊢ T0 ➡[h] T1 →
+   ∀V2. ⦃G, L0⦄ ⊢ V0 ➡[h] V2 → ∀U2. ⬆*[1] V2 ≡ U2 →
+   ∀W2. ⦃G, L0⦄ ⊢ W0 ➡[h] W2 → ∀T2. ⦃G, L0.ⓓW0⦄ ⊢ T0 ➡[h] T2 →
+   ∀L1. ⦃G, L0⦄ ⊢ ➡[h, ⓐV0.ⓓ{p}W0.T0] L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡[h, ⓐV0.ⓓ{p}W0.T0] L2 →
+   ∃∃T. ⦃G, L1⦄ ⊢ ⓓ{p}W1.ⓐU1.T1 ➡[h] T & ⦃G, L2⦄ ⊢ ⓓ{p}W2.ⓐU2.T2 ➡[h] T.
+#h #p #G #L0 #V0 #W0 #T0 #IH #V1 #HV01 #U1 #HVU1 #W1 #HW01 #T1 #HT01
+#V2 #HV02 #U2 #HVU2 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
+elim (lfpr_inv_flat … HL01) -HL01 #H1V0 #HL01
+elim (lfpr_inv_bind … HL01) -HL01 #H1W0 #H1T0
+elim (lfpr_inv_flat … HL02) -HL02 #H2V0 #HL02
+elim (lfpr_inv_bind … HL02) -HL02 #H2W0 #H2T0
+elim (IH … HV01 … HV02 … H1V0 … H2V0) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
+elim (IH … HW01 … HW02 … H1W0 … H2W0) /2 width=1 by/
+elim (IH … HT01 … HT02 (L1.ⓓW1) … (L2.ⓓW2)) /2 width=4 by lfpr_pair_repl_dx/ -L0 -V0 -W0 -T0
+elim (cpm_lifts … HV1 … (L1.ⓓW1) … HVU1) -HVU1 /3 width=2 by drops_refl, drops_drop/ #U #HVU
+elim (cpm_lifts … HV2 … (L2.ⓓW2) … HVU2) -HVU2 /3 width=2 by drops_refl, drops_drop/ #X #HX
+lapply (lifts_mono … HX … HVU) -HX #H destruct
+/4 width=7 by cpm_bind, cpr_flat, ex2_intro/ (**) (* full auto not tried *)
+qed-.
+
+theorem cpr_conf_lfpr: ∀h,G. lfxs_confluent (cpm 0 h G) (cpm 0 h G) (cpm 0 h G) (cpm 0 h G).
+#h #G #L0 #T0 @(fqup_wf_ind_eq … G L0 T0) -G -L0 -T0 #G #L #T #IH #G0 #L0 * [| * ]
+[ #I0 #HG #HL #HT #T1 #H1 #T2 #H2 #L1 #HL01 #L2 #HL02 destruct
+  elim (cpr_inv_atom1_drops … H1) -H1
+  elim (cpr_inv_atom1_drops … H2) -H2
+  [ #H2 #H1 destruct
+    /2 width=1 by cpr_conf_lfpr_atom_atom/
+  | * #K0 #V0 #V2 #i2 #HLK0 #HV02 #HVT2 #H2 #H1 destruct
+    /3 width=10 by cpr_conf_lfpr_atom_delta/
+  | #H2 * #K0 #V0 #V1 #i1 #HLK0 #HV01 #HVT1 #H1 destruct
+    /4 width=10 by ex2_commute, cpr_conf_lfpr_atom_delta/
+  | * #X #Y #V2 #z #H #HV02 #HVT2 #H2
+    * #K0 #V0 #V1 #i #HLK0 #HV01 #HVT1 #H1 destruct
+    /3 width=17 by cpr_conf_lfpr_delta_delta/
+  ]
+| #p #I #V0 #T0 #HG #HL #HT #X1 #H1 #X2 #H2 #L1 #HL01 #L2 #HL02 destruct
+  elim (cpm_inv_bind1 … H1) -H1 *
+  [ #V1 #T1 #HV01 #HT01 #H1
+  | #T1 #HT01 #HXT1 #H11 #H12
+  ]
+  elim (cpm_inv_bind1 … H2) -H2 *
+  [1,3: #V2 #T2 #HV02 #HT02 #H2
+  |2,4: #T2 #HT02 #HXT2 #H21 #H22
+  ] destruct
+  [ /3 width=10 by cpr_conf_lfpr_bind_bind/
+  | /4 width=11 by ex2_commute, cpr_conf_lfpr_bind_zeta/
+  | /3 width=11 by cpr_conf_lfpr_bind_zeta/
+  | /3 width=12 by cpr_conf_lfpr_zeta_zeta/
+  ]
+| #I #V0 #T0 #HG #HL #HT #X1 #H1 #X2 #H2 #L1 #HL01 #L2 #HL02 destruct
+  elim (cpr_inv_flat1 … H1) -H1 *
+  [ #V1 #T1 #HV01 #HT01 #H1
+  | #HX1 #H1
+  | #p1 #V1 #Y1 #W1 #Z1 #T1 #HV01 #HYW1 #HZT1 #H11 #H12 #H13
+  | #p1 #V1 #U1 #Y1 #W1 #Z1 #T1 #HV01 #HVU1 #HYW1 #HZT1 #H11 #H12 #H13
+  ]
+  elim (cpr_inv_flat1 … H2) -H2 *
+  [1,5,9,13: #V2 #T2 #HV02 #HT02 #H2
+  |2,6,10,14: #HX2 #H2
+  |3,7,11,15: #p2 #V2 #Y2 #W2 #Z2 #T2 #HV02 #HYW2 #HZT2 #H21 #H22 #H23
+  |4,8,12,16: #p2 #V2 #U2 #Y2 #W2 #Z2 #T2 #HV02 #HVU2 #HYW2 #HZT2 #H21 #H22 #H23
+  ] destruct
+  [ /3 width=10 by cpr_conf_lfpr_flat_flat/
+  | /4 width=8 by ex2_commute, cpr_conf_lfpr_flat_epsilon/
+  | /4 width=12 by ex2_commute, cpr_conf_lfpr_flat_beta/
+  | /4 width=14 by ex2_commute, cpr_conf_lfpr_flat_theta/
+  | /3 width=8 by cpr_conf_lfpr_flat_epsilon/
+  | /3 width=8 by cpr_conf_lfpr_epsilon_epsilon/
+  | /3 width=12 by cpr_conf_lfpr_flat_beta/
+  | /3 width=13 by cpr_conf_lfpr_beta_beta/
+  | /3 width=14 by cpr_conf_lfpr_flat_theta/
+  | /3 width=17 by cpr_conf_lfpr_theta_theta/
+  ]
+]
+qed-.
+
+(* Basic_1: includes: pr0_confluence pr2_confluence *)
+theorem cpr_conf: ∀h,G,L. confluent … (cpm 0 h G L).
+/2 width=6 by cpr_conf_lfpr/ qed-.
+
+(* Properties with context-sensitive parallel r-transition for terms ********)
+
+lemma lfpr_cpr_conf_dx: ∀h,G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡[h] T1 → ∀L1. ⦃G, L0⦄ ⊢ ➡[h, T0] L1 →
+                        ∃∃T. ⦃G, L1⦄ ⊢ T0 ➡[h] T & ⦃G, L1⦄ ⊢ T1 ➡[h] T.
+#h #G #L0 #T0 #T1 #HT01 #L1 #HL01
+elim (cpr_conf_lfpr … HT01 T0 … HL01 … HL01) /2 width=3 by ex2_intro/
+qed-.
+
+lemma lfpr_cpr_conf_sn: ∀h,G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡[h] T1 → ∀L1. ⦃G, L0⦄ ⊢ ➡[h, T0] L1 →
+                        ∃∃T. ⦃G, L1⦄ ⊢ T0 ➡[h] T & ⦃G, L0⦄ ⊢ T1 ➡[h] T.
+#h #G #L0 #T0 #T1 #HT01 #L1 #HL01
+elim (cpr_conf_lfpr … HT01 T0 … L0 … HL01) /2 width=3 by ex2_intro/
+qed-.
+
+(* Main properties **********************************************************)
+
+(*
+
+theorem lpr_conf: ∀G. confluent … (lpr G).
+/3 width=6 by lpx_sn_conf, cpr_conf_lpr/
+qed-.
+
+*)
index 57c0184beedb536c36ad05b3ccc961fdff564be9..c0ca76840f32238d9fdeb0438d47313eab8333d7 100644 (file)
@@ -46,6 +46,12 @@ lemma lfpx_gref: ∀h,I,G,L1,L2,V1,V2,l.
                  ⦃G, L1⦄ ⊢ ⬈[h, §l] L2 → ⦃G, L1.ⓑ{I}V1⦄ ⊢ ⬈[h, §l] L2.ⓑ{I}V2.
 /2 width=1 by lfxs_gref/ qed.
 
+lemma lfpx_pair_repl_dx: ∀h,I,G,L1,L2,T,V,V1.
+                         ⦃G, L1.ⓑ{I}V⦄ ⊢ ⬈[h, T] L2.ⓑ{I}V1 →
+                         ∀V2. ⦃G, L1⦄ ⊢ V ⬈[h] V2 →
+                         ⦃G, L1.ⓑ{I}V⦄ ⊢ ⬈[h, T] L2.ⓑ{I}V2.
+/2 width=2 by lfxs_pair_repl_dx/ qed-.
+
 (* Basic inversion lemmas ***************************************************)
 
 lemma lfpx_inv_atom_sn: ∀h,I,G,Y2. ⦃G, ⋆⦄ ⊢ ⬈[h, ⓪{I}] Y2 → Y2 = ⋆.
index 55c4051628ca455bd10df0706428d3b9e475172e..086e71e87bd791fd522681adf3feb065a4419220 100644 (file)
@@ -30,3 +30,11 @@ lemma lfpx_drops_conf: ∀h,G. dropable_sn (cpx h G).
 
 lemma lfpx_drops_trans: ∀h,G. dropable_dx (cpx h G).
 /2 width=5 by lfxs_dropable_dx/ qed-.
+
+lemma lfpx_inv_lref_sn: ∀h,G,L1,L2,i. ⦃G, L1⦄ ⊢ ⬈[h, #i] L2 → ∀I,K1,V1. ⬇*[i] L1 ≡ K1.ⓑ{I}V1 →
+                        ∃∃K2,V2. ⬇*[i] L2 ≡ K2.ⓑ{I}V2 & ⦃G, K1⦄ ⊢ ⬈[h, V1] K2 & ⦃G, K1⦄ ⊢ V1 ⬈[h] V2.
+/2 width=3 by lfxs_inv_lref_sn/ qed-.
+
+lemma lfpx_inv_lref_dx: ∀h,G,L1,L2,i. ⦃G, L1⦄ ⊢ ⬈[h, #i] L2 → ∀I,K2,V2. ⬇*[i] L2 ≡ K2.ⓑ{I}V2 →
+                        ∃∃K1,V1. ⬇*[i] L1 ≡ K1.ⓑ{I}V1 & ⦃G, K1⦄ ⊢ ⬈[h, V1] K2 & ⦃G, K1⦄ ⊢ V1 ⬈[h] V2.
+/2 width=3 by lfxs_inv_lref_dx/ qed-.
diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lpr_lpr.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lpr_lpr.ma
deleted file mode 100644 (file)
index fad8b48..0000000
+++ /dev/null
@@ -1,357 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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 "basic_2/substitution/lpx_sn_lpx_sn.ma".
-include "basic_2/multiple/fqup.ma".
-include "basic_2/reduction/lpr_drop.ma".
-
-(* SN PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS *****************************)
-
-(* Main properties on context-sensitive parallel reduction for terms ********)
-
-fact cpr_conf_lpr_atom_atom:
-   ∀I,G,L1,L2. ∃∃T. ⦃G, L1⦄ ⊢ ⓪{I} ➡ T & ⦃G, L2⦄ ⊢ ⓪{I} ➡ T.
-/2 width=3 by cpr_atom, ex2_intro/ qed-.
-
-fact cpr_conf_lpr_atom_delta:
-   ∀G,L0,i. (
-      ∀L,T. ⦃G, L0, #i⦄ ⊐+ ⦃G, L, T⦄ →
-      ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
-      ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
-      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
-   ) →
-   ∀K0,V0. ⬇[i] L0 ≡ K0.ⓓV0 →
-   ∀V2. ⦃G, K0⦄ ⊢ V0 ➡ V2 → ∀T2. ⬆[O, i + 1] V2 ≡ T2 →
-   ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
-   ∃∃T. ⦃G, L1⦄ ⊢ #i ➡ T & ⦃G, L2⦄ ⊢ T2 ➡ T.
-#G #L0 #i #IH #K0 #V0 #HLK0 #V2 #HV02 #T2 #HVT2 #L1 #HL01 #L2 #HL02
-elim (lpr_drop_conf … HLK0 … HL01) -HL01 #X1 #H1 #HLK1
-elim (lpr_inv_pair1 … H1) -H1 #K1 #V1 #HK01 #HV01 #H destruct
-elim (lpr_drop_conf … HLK0 … HL02) -HL02 #X2 #H2 #HLK2
-elim (lpr_inv_pair1 … H2) -H2 #K2 #W2 #HK02 #_ #H destruct
-lapply (drop_fwd_drop2 … HLK2) -W2 #HLK2
-lapply (fqup_lref … G … HLK0) -HLK0 #HLK0
-elim (IH … HLK0 … HV01 … HV02 … HK01 … HK02) -L0 -K0 -V0 #V #HV1 #HV2
-elim (lift_total V 0 (i+1))
-/3 width=12 by cpr_lift, cpr_delta, ex2_intro/
-qed-.
-
-(* Basic_1: includes: pr0_delta_delta pr2_delta_delta *)
-fact cpr_conf_lpr_delta_delta:
-   ∀G,L0,i. (
-      ∀L,T. ⦃G, L0, #i⦄ ⊐+ ⦃G, L, T⦄ →
-      ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
-      ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
-      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
-   ) →
-   ∀K0,V0. ⬇[i] L0 ≡ K0.ⓓV0 →
-   ∀V1. ⦃G, K0⦄ ⊢ V0 ➡ V1 → ∀T1. ⬆[O, i + 1] V1 ≡ T1 →
-   ∀KX,VX. ⬇[i] L0 ≡ KX.ⓓVX →
-   ∀V2. ⦃G, KX⦄ ⊢ VX ➡ V2 → ∀T2. ⬆[O, i + 1] V2 ≡ T2 →
-   ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
-   ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡ T & ⦃G, L2⦄ ⊢ T2 ➡ T.
-#G #L0 #i #IH #K0 #V0 #HLK0 #V1 #HV01 #T1 #HVT1
-#KX #VX #H #V2 #HV02 #T2 #HVT2 #L1 #HL01 #L2 #HL02
-lapply (drop_mono … H … HLK0) -H #H destruct
-elim (lpr_drop_conf … HLK0 … HL01) -HL01 #X1 #H1 #HLK1
-elim (lpr_inv_pair1 … H1) -H1 #K1 #W1 #HK01 #_ #H destruct
-lapply (drop_fwd_drop2 … HLK1) -W1 #HLK1
-elim (lpr_drop_conf … HLK0 … HL02) -HL02 #X2 #H2 #HLK2
-elim (lpr_inv_pair1 … H2) -H2 #K2 #W2 #HK02 #_ #H destruct
-lapply (drop_fwd_drop2 … HLK2) -W2 #HLK2
-lapply (fqup_lref … G … HLK0) -HLK0 #HLK0
-elim (IH … HLK0 … HV01 … HV02 … HK01 … HK02) -L0 -K0 -V0 #V #HV1 #HV2
-elim (lift_total V 0 (i+1)) /3 width=12 by cpr_lift, ex2_intro/
-qed-.
-
-fact cpr_conf_lpr_bind_bind:
-   ∀a,I,G,L0,V0,T0. (
-      ∀L,T. ⦃G, L0, ⓑ{a,I}V0.T0⦄ ⊐+ ⦃G, L, T⦄ →
-      ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
-      ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
-      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
-   ) →
-   ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀T1. ⦃G, L0.ⓑ{I}V0⦄ ⊢ T0 ➡ T1 →
-   ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀T2. ⦃G, L0.ⓑ{I}V0⦄ ⊢ T0 ➡ T2 →
-   ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
-   ∃∃T. ⦃G, L1⦄ ⊢ ⓑ{a,I}V1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓑ{a,I}V2.T2 ➡ T.
-#a #I #G #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01
-#V2 #HV02 #T2 #HT02 #L1 #HL01 #L2 #HL02
-elim (IH … HV01 … HV02 … HL01 … HL02) //
-elim (IH … HT01 … HT02 (L1.ⓑ{I}V1) … (L2.ⓑ{I}V2)) -IH
-/3 width=5 by lpr_pair, cpr_bind, ex2_intro/
-qed-.
-
-fact cpr_conf_lpr_bind_zeta:
-   ∀G,L0,V0,T0. (
-      ∀L,T. ⦃G, L0, +ⓓV0.T0⦄ ⊐+ ⦃G, L, T⦄ →
-      ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
-      ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
-      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
-   ) →
-   ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀T1. ⦃G, L0.ⓓV0⦄ ⊢ T0 ➡ T1 →
-   ∀T2. ⦃G, L0.ⓓV0⦄ ⊢ T0 ➡ T2 → ∀X2. ⬆[O, 1] X2 ≡ T2 →
-   ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
-   ∃∃T. ⦃G, L1⦄ ⊢ +ⓓV1.T1 ➡ T & ⦃G, L2⦄ ⊢ X2 ➡ T.
-#G #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01
-#T2 #HT02 #X2 #HXT2 #L1 #HL01 #L2 #HL02
-elim (IH … HT01 … HT02 (L1.ⓓV1) … (L2.ⓓV1)) -IH -HT01 -HT02 /2 width=1 by lpr_pair/ -L0 -V0 -T0 #T #HT1 #HT2
-elim (cpr_inv_lift1 … HT2 L2 … HXT2) -T2 /3 width=3 by cpr_zeta, drop_drop, ex2_intro/
-qed-.
-
-fact cpr_conf_lpr_zeta_zeta:
-   ∀G,L0,V0,T0. (
-      ∀L,T. ⦃G, L0, +ⓓV0.T0⦄ ⊐+ ⦃G, L, T⦄ →
-      ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
-      ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
-      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
-   ) →
-   ∀T1. ⦃G, L0.ⓓV0⦄ ⊢ T0 ➡ T1 → ∀X1. ⬆[O, 1] X1 ≡ T1 →
-   ∀T2. ⦃G, L0.ⓓV0⦄ ⊢ T0 ➡ T2 → ∀X2. ⬆[O, 1] X2 ≡ T2 →
-   ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
-   ∃∃T. ⦃G, L1⦄ ⊢ X1 ➡ T & ⦃G, L2⦄ ⊢ X2 ➡ T.
-#G #L0 #V0 #T0 #IH #T1 #HT01 #X1 #HXT1
-#T2 #HT02 #X2 #HXT2 #L1 #HL01 #L2 #HL02
-elim (IH … HT01 … HT02 (L1.ⓓV0) … (L2.ⓓV0)) -IH -HT01 -HT02 /2 width=1 by lpr_pair/ -L0 -T0 #T #HT1 #HT2
-elim (cpr_inv_lift1 … HT1 L1 … HXT1) -T1 /2 width=2 by drop_drop/ #T1 #HT1 #HXT1
-elim (cpr_inv_lift1 … HT2 L2 … HXT2) -T2 /2 width=2 by drop_drop/ #T2 #HT2 #HXT2
-lapply (lift_inj … HT2 … HT1) -T #H destruct /2 width=3 by ex2_intro/
-qed-.
-
-fact cpr_conf_lpr_flat_flat:
-   ∀I,G,L0,V0,T0. (
-      ∀L,T. ⦃G, L0, ⓕ{I}V0.T0⦄ ⊐+ ⦃G, L, T⦄ →
-      ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
-      ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
-      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
-   ) →
-   ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 →
-   ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀T2. ⦃G, L0⦄ ⊢ T0 ➡ T2 →
-   ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
-   ∃∃T. ⦃G, L1⦄ ⊢ ⓕ{I}V1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓕ{I}V2.T2 ➡ T.
-#I #G #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01
-#V2 #HV02 #T2 #HT02 #L1 #HL01 #L2 #HL02
-elim (IH … HV01 … HV02 … HL01 … HL02) //
-elim (IH … HT01 … HT02 … HL01 … HL02) /3 width=5 by cpr_flat, ex2_intro/
-qed-.
-
-fact cpr_conf_lpr_flat_eps:
-   ∀G,L0,V0,T0. (
-      ∀L,T. ⦃G, L0, ⓝV0.T0⦄ ⊐+ ⦃G, L, T⦄ →
-      ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
-      ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
-      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
-   ) →
-   ∀V1,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 → ∀T2. ⦃G, L0⦄ ⊢ T0 ➡ T2 →
-   ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
-   ∃∃T. ⦃G, L1⦄ ⊢ ⓝV1.T1 ➡ T & ⦃G, L2⦄ ⊢ T2 ➡ T.
-#G #L0 #V0 #T0 #IH #V1 #T1 #HT01
-#T2 #HT02 #L1 #HL01 #L2 #HL02
-elim (IH … HT01 … HT02 … HL01 … HL02) // -L0 -V0 -T0 /3 width=3 by cpr_eps, ex2_intro/
-qed-.
-
-fact cpr_conf_lpr_eps_eps:
-   ∀G,L0,V0,T0. (
-      ∀L,T. ⦃G, L0, ⓝV0.T0⦄ ⊐+ ⦃G, L, T⦄ →
-      ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
-      ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
-      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
-   ) →
-   ∀T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 → ∀T2. ⦃G, L0⦄ ⊢ T0 ➡ T2 →
-   ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
-   ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡ T & ⦃G, L2⦄ ⊢ T2 ➡ T.
-#G #L0 #V0 #T0 #IH #T1 #HT01
-#T2 #HT02 #L1 #HL01 #L2 #HL02
-elim (IH … HT01 … HT02 … HL01 … HL02) // -L0 -V0 -T0 /2 width=3 by ex2_intro/
-qed-.
-
-fact cpr_conf_lpr_flat_beta:
-   ∀a,G,L0,V0,W0,T0. (
-      ∀L,T. ⦃G, L0, ⓐV0.ⓛ{a}W0.T0⦄ ⊐+ ⦃G, L, T⦄ →
-      ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
-      ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
-      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
-   ) →
-   ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀T1. ⦃G, L0⦄ ⊢ ⓛ{a}W0.T0 ➡ T1 →
-   ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀W2. ⦃G, L0⦄ ⊢ W0 ➡ W2 → ∀T2. ⦃G, L0.ⓛW0⦄ ⊢ T0 ➡ T2 →
-   ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
-   ∃∃T. ⦃G, L1⦄ ⊢ ⓐV1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓓ{a}ⓝW2.V2.T2 ➡ T.
-#a #G #L0 #V0 #W0 #T0 #IH #V1 #HV01 #X #H
-#V2 #HV02 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
-elim (cpr_inv_abst1 … H) -H #W1 #T1 #HW01 #HT01 #H destruct
-elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
-elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1 by/ #W #HW1 #HW2
-elim (IH … HT01 … HT02 (L1.ⓛW1) … (L2.ⓛW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0 #T #HT1 #HT2
-lapply (lsubr_cpr_trans … HT2 (L2.ⓓⓝW2.V2) ?) -HT2 /2 width=1 by lsubr_beta/ (**) (* full auto not tried *)
-/4 width=5 by cpr_bind, cpr_flat, cpr_beta, ex2_intro/
-qed-.
-
-(* Basic-1: includes:
-            pr0_cong_upsilon_refl pr0_cong_upsilon_zeta
-            pr0_cong_upsilon_cong pr0_cong_upsilon_delta
-*)
-fact cpr_conf_lpr_flat_theta:
-   ∀a,G,L0,V0,W0,T0. (
-      ∀L,T. ⦃G, L0, ⓐV0.ⓓ{a}W0.T0⦄ ⊐+ ⦃G, L, T⦄ →
-      ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
-      ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
-      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
-   ) →
-   ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀T1. ⦃G, L0⦄ ⊢ ⓓ{a}W0.T0 ➡ T1 →
-   ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀U2. ⬆[O, 1] V2 ≡ U2 →
-   ∀W2. ⦃G, L0⦄ ⊢ W0 ➡ W2 → ∀T2. ⦃G, L0.ⓓW0⦄ ⊢ T0 ➡ T2 →
-   ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
-   ∃∃T. ⦃G, L1⦄ ⊢ ⓐV1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓓ{a}W2.ⓐU2.T2 ➡ T.
-#a #G #L0 #V0 #W0 #T0 #IH #V1 #HV01 #X #H
-#V2 #HV02 #U2 #HVU2 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
-elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
-elim (lift_total V 0 1) #U #HVU
-lapply (cpr_lift … HV2 (L2.ⓓW2) … HVU2 … HVU) -HVU2 /2 width=2 by drop_drop/ #HU2
-elim (cpr_inv_abbr1 … H) -H *
-[ #W1 #T1 #HW01 #HT01 #H destruct
-  elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1 by/
-  elim (IH … HT01 … HT02 (L1.ⓓW1) … (L2.ⓓW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0
-  /4 width=7 by cpr_bind, cpr_flat, cpr_theta, ex2_intro/
-| #T1 #HT01 #HXT1 #H destruct
-  elim (IH … HT01 … HT02 (L1.ⓓW2) … (L2.ⓓW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0 #T #HT1 #HT2
-  elim (cpr_inv_lift1 … HT1 L1 … HXT1) -HXT1
-  /4 width=9 by cpr_flat, cpr_zeta, drop_drop, lift_flat, ex2_intro/
-]
-qed-.
-
-fact cpr_conf_lpr_beta_beta:
-   ∀a,G,L0,V0,W0,T0. (
-      ∀L,T. ⦃G, L0, ⓐV0.ⓛ{a}W0.T0⦄ ⊐+ ⦃G, L, T⦄ →
-      ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
-      ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
-      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
-   ) →
-   ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀W1. ⦃G, L0⦄ ⊢ W0 ➡ W1 → ∀T1. ⦃G, L0.ⓛW0⦄ ⊢ T0 ➡ T1 →
-   ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀W2. ⦃G, L0⦄ ⊢ W0 ➡ W2 → ∀T2. ⦃G, L0.ⓛW0⦄ ⊢ T0 ➡ T2 →
-   ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
-   ∃∃T. ⦃G, L1⦄ ⊢ ⓓ{a}ⓝW1.V1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓓ{a}ⓝW2.V2.T2 ➡ T.
-#a #G #L0 #V0 #W0 #T0 #IH #V1 #HV01 #W1 #HW01 #T1 #HT01
-#V2 #HV02 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
-elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
-elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1 by/ #W #HW1 #HW2
-elim (IH … HT01 … HT02 (L1.ⓛW1) … (L2.ⓛW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0 #T #HT1 #HT2
-lapply (lsubr_cpr_trans … HT1 (L1.ⓓⓝW1.V1) ?) -HT1 /2 width=1 by lsubr_beta/
-lapply (lsubr_cpr_trans … HT2 (L2.ⓓⓝW2.V2) ?) -HT2 /2 width=1 by lsubr_beta/
-/4 width=5 by cpr_bind, cpr_flat, ex2_intro/ (**) (* full auto not tried *)
-qed-.
-
-(* Basic_1: was: pr0_upsilon_upsilon *)
-fact cpr_conf_lpr_theta_theta:
-   ∀a,G,L0,V0,W0,T0. (
-      ∀L,T. ⦃G, L0, ⓐV0.ⓓ{a}W0.T0⦄ ⊐+ ⦃G, L, T⦄ →
-      ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 →
-      ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 →
-      ∃∃T0. ⦃G, L1⦄ ⊢ T1 ➡ T0 & ⦃G, L2⦄ ⊢ T2 ➡ T0
-   ) →
-   ∀V1. ⦃G, L0⦄ ⊢ V0 ➡ V1 → ∀U1. ⬆[O, 1] V1 ≡ U1 →
-   ∀W1. ⦃G, L0⦄ ⊢ W0 ➡ W1 → ∀T1. ⦃G, L0.ⓓW0⦄ ⊢ T0 ➡ T1 →
-   ∀V2. ⦃G, L0⦄ ⊢ V0 ➡ V2 → ∀U2. ⬆[O, 1] V2 ≡ U2 →
-   ∀W2. ⦃G, L0⦄ ⊢ W0 ➡ W2 → ∀T2. ⦃G, L0.ⓓW0⦄ ⊢ T0 ➡ T2 →
-   ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L0⦄ ⊢ ➡ L2 →
-   ∃∃T. ⦃G, L1⦄ ⊢ ⓓ{a}W1.ⓐU1.T1 ➡ T & ⦃G, L2⦄ ⊢ ⓓ{a}W2.ⓐU2.T2 ➡ T.
-#a #G #L0 #V0 #W0 #T0 #IH #V1 #HV01 #U1 #HVU1 #W1 #HW01 #T1 #HT01
-#V2 #HV02 #U2 #HVU2 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02
-elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1 by/ #V #HV1 #HV2
-elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1 by/
-elim (IH … HT01 … HT02 (L1.ⓓW1) … (L2.ⓓW2)) /2 width=1 by lpr_pair/ -L0 -V0 -W0 -T0
-elim (lift_total V 0 1) #U #HVU
-lapply (cpr_lift … HV1 (L1.ⓓW1) … HVU1 … HVU) -HVU1 /2 width=2 by drop_drop/
-lapply (cpr_lift … HV2 (L2.ⓓW2) … HVU2 … HVU) -HVU2 /2 width=2 by drop_drop/
-/4 width=7 by cpr_bind, cpr_flat, ex2_intro/ (**) (* full auto not tried *)
-qed-.
-
-theorem cpr_conf_lpr: ∀G. lpx_sn_confluent (cpr G) (cpr G).
-#G #L0 #T0 @(fqup_wf_ind_eq … G L0 T0) -G -L0 -T0 #G #L #T #IH #G0 #L0 * [| * ]
-[ #I0 #HG #HL #HT #T1 #H1 #T2 #H2 #L1 #HL01 #L2 #HL02 destruct
-  elim (cpr_inv_atom1 … H1) -H1
-  elim (cpr_inv_atom1 … H2) -H2
-  [ #H2 #H1 destruct
-    /2 width=1 by cpr_conf_lpr_atom_atom/
-  | * #K0 #V0 #V2 #i2 #HLK0 #HV02 #HVT2 #H2 #H1 destruct
-    /3 width=10 by cpr_conf_lpr_atom_delta/
-  | #H2 * #K0 #V0 #V1 #i1 #HLK0 #HV01 #HVT1 #H1 destruct
-    /4 width=10 by ex2_commute, cpr_conf_lpr_atom_delta/
-  | * #X #Y #V2 #z #H #HV02 #HVT2 #H2
-    * #K0 #V0 #V1 #i #HLK0 #HV01 #HVT1 #H1 destruct
-    /3 width=17 by cpr_conf_lpr_delta_delta/
-  ]
-| #a #I #V0 #T0 #HG #HL #HT #X1 #H1 #X2 #H2 #L1 #HL01 #L2 #HL02 destruct
-  elim (cpr_inv_bind1 … H1) -H1 *
-  [ #V1 #T1 #HV01 #HT01 #H1
-  | #T1 #HT01 #HXT1 #H11 #H12
-  ]
-  elim (cpr_inv_bind1 … H2) -H2 *
-  [1,3: #V2 #T2 #HV02 #HT02 #H2
-  |2,4: #T2 #HT02 #HXT2 #H21 #H22
-  ] destruct
-  [ /3 width=10 by cpr_conf_lpr_bind_bind/
-  | /4 width=11 by ex2_commute, cpr_conf_lpr_bind_zeta/
-  | /3 width=11 by cpr_conf_lpr_bind_zeta/
-  | /3 width=12 by cpr_conf_lpr_zeta_zeta/
-  ]
-| #I #V0 #T0 #HG #HL #HT #X1 #H1 #X2 #H2 #L1 #HL01 #L2 #HL02 destruct
-  elim (cpr_inv_flat1 … H1) -H1 *
-  [ #V1 #T1 #HV01 #HT01 #H1
-  | #HX1 #H1
-  | #a1 #V1 #Y1 #W1 #Z1 #T1 #HV01 #HYW1 #HZT1 #H11 #H12 #H13
-  | #a1 #V1 #U1 #Y1 #W1 #Z1 #T1 #HV01 #HVU1 #HYW1 #HZT1 #H11 #H12 #H13
-  ]
-  elim (cpr_inv_flat1 … H2) -H2 *
-  [1,5,9,13: #V2 #T2 #HV02 #HT02 #H2
-  |2,6,10,14: #HX2 #H2
-  |3,7,11,15: #a2 #V2 #Y2 #W2 #Z2 #T2 #HV02 #HYW2 #HZT2 #H21 #H22 #H23
-  |4,8,12,16: #a2 #V2 #U2 #Y2 #W2 #Z2 #T2 #HV02 #HVU2 #HYW2 #HZT2 #H21 #H22 #H23
-  ] destruct
-  [ /3 width=10 by cpr_conf_lpr_flat_flat/
-  | /4 width=8 by ex2_commute, cpr_conf_lpr_flat_eps/
-  | /4 width=12 by ex2_commute, cpr_conf_lpr_flat_beta/
-  | /4 width=14 by ex2_commute, cpr_conf_lpr_flat_theta/
-  | /3 width=8 by cpr_conf_lpr_flat_eps/
-  | /3 width=7 by cpr_conf_lpr_eps_eps/
-  | /3 width=12 by cpr_conf_lpr_flat_beta/
-  | /3 width=13 by cpr_conf_lpr_beta_beta/
-  | /3 width=14 by cpr_conf_lpr_flat_theta/
-  | /3 width=17 by cpr_conf_lpr_theta_theta/
-  ]
-]
-qed-.
-
-(* Basic_1: includes: pr0_confluence pr2_confluence *)
-theorem cpr_conf: ∀G,L. confluent … (cpr G L).
-/2 width=6 by cpr_conf_lpr/ qed-.
-
-(* Properties on context-sensitive parallel reduction for terms *************)
-
-lemma lpr_cpr_conf_dx: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 → ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 →
-                       ∃∃T. ⦃G, L1⦄ ⊢ T0 ➡ T & ⦃G, L1⦄ ⊢ T1 ➡ T.
-#G #L0 #T0 #T1 #HT01 #L1 #HL01
-elim (cpr_conf_lpr … HT01 T0 … HL01 … HL01) /2 width=3 by ex2_intro/
-qed-.
-
-lemma lpr_cpr_conf_sn: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 → ∀L1. ⦃G, L0⦄ ⊢ ➡ L1 →
-                       ∃∃T. ⦃G, L1⦄ ⊢ T0 ➡ T & ⦃G, L0⦄ ⊢ T1 ➡ T.
-#G #L0 #T0 #T1 #HT01 #L1 #HL01
-elim (cpr_conf_lpr … HT01 T0 … L0 … HL01) /2 width=3 by ex2_intro/
-qed-.
-
-(* Main properties **********************************************************)
-
-theorem lpr_conf: ∀G. confluent … (lpr G).
-/3 width=6 by lpx_sn_conf, cpr_conf_lpr/
-qed-.
index 63773c485f166e9de5a3bb7f7ce9bce3def106a7..d0eb3451c77b9b030969eddbb7429aab36139f3b 100644 (file)
@@ -3,4 +3,4 @@ cpx.ma cpx_simple.ma cpx_drops.ma cpx_lsubr.ma
 lfpx.ma lfpx_length.ma lfpx_drops.ma lfpx_fqup.ma 
 cpm.ma cpm_simple.ma cpm_drops.ma cpm_lsubr.ma cpm_cpx.ma
 cpr.ma cpr_drops.ma
-lfpr.ma lfpr_length.ma lfpr_drops.ma lfpr_fqup.ma lfpr_lfpx.ma
+lfpr.ma lfpr_length.ma lfpr_drops.ma lfpr_fqup.ma lfpr_lfpx.ma lfpr_lfpr.ma
index b3c7a7fe06e8d24c8164bf582743c0ad46004757..88a8bd542182ff72e19c288ded5cbe5915d3a725 100644 (file)
@@ -26,6 +26,13 @@ definition lfxs (R) (T): relation lenv ≝
 interpretation "generic extension on referred entries (local environment)"
    'RelationStar R T L1 L2 = (lfxs R T L1 L2).
 
+definition lfxs_confluent: relation4 (relation3 lenv term term)
+                                     (relation3 lenv term term) … ≝
+                           λR1,R2,RP1,RP2.
+                           ∀L0,T0,T1. R1 L0 T0 T1 → ∀T2. R2 L0 T0 T2 →
+                           ∀L1. L0 ⦻*[RP1, T0] L1 → ∀L2. L0 ⦻*[RP2, T0] L2 →
+                           ∃∃T. R2 L1 T1 T & R1 L2 T2 T.
+
 (* Basic properties ***********************************************************)
 
 lemma lfxs_atom: ∀R,I. ⋆ ⦻*[R, ⓪{I}] ⋆.
@@ -52,6 +59,14 @@ lemma lfxs_gref: ∀R,I,L1,L2,V1,V2,l.
 #R #I #L1 #L2 #V1 #V2 #l * /3 width=3 by lexs_push, frees_gref, ex2_intro/
 qed.
 
+lemma lfxs_pair_repl_dx: ∀R,I,L1,L2,T,V,V1.
+                         L1.ⓑ{I}V ⦻*[R, T] L2.ⓑ{I}V1 →
+                         ∀V2. R L1 V V2 →
+                         L1.ⓑ{I}V ⦻*[R, T] L2.ⓑ{I}V2.
+#R #I #L1 #L2 #T #V #V1 * #f #Hf #HL12 #V2 #HR
+/3 width=5 by lexs_pair_repl, ex2_intro/
+qed-.
+
 lemma lfxs_co: ∀R1,R2. (∀L,T1,T2. R1 L T1 T2 → R2 L T1 T2) →
                ∀L1,L2,T. L1 ⦻*[R1, T] L2 → L1 ⦻*[R2, T] L2.
 #R1 #R2 #HR #L1 #L2 #T * /4 width=7 by lexs_co, ex2_intro/
index 0a64c396bb4903ff10b878985b7c53a9decf9caa..f43eb4e95956f239b843d1aa744e5a955455bff0 100644 (file)
@@ -78,3 +78,17 @@ lemma lfxs_inv_lift_bi: ∀R,L1,L2,U. L1 ⦻*[R, U] L2 →
 elim (lfxs_dropable_sn … HLK1 … HL12 … HTU) -L1 -U // #Y #HK12 #HY
 lapply (drops_mono … HY … HLK2) -L2 -i #H destruct //
 qed-.
+
+lemma lfxs_inv_lref_sn: ∀R,L1,L2,i. L1 ⦻*[R, #i] L2 → ∀I,K1,V1. ⬇*[i] L1 ≡ K1.ⓑ{I}V1 →
+                        ∃∃K2,V2. ⬇*[i] L2 ≡ K2.ⓑ{I}V2 & K1 ⦻*[R, V1] K2 & R K1 V1 V2.
+#R #L1 #L2 #i #HL12 #I #K1 #V1 #HLK1 elim (lfxs_dropable_sn … HLK1 … HL12 (#0)) -HLK1 -HL12 //
+#Y #HY #HLK2 elim (lfxs_inv_zero_pair_sn … HY) -HY
+#K2 #V2 #HK12 #HV12 #H destruct /2 width=5 by ex3_2_intro/
+qed-.
+
+lemma lfxs_inv_lref_dx: ∀R,L1,L2,i. L1 ⦻*[R, #i] L2 → ∀I,K2,V2. ⬇*[i] L2 ≡ K2.ⓑ{I}V2 →
+                        ∃∃K1,V1. ⬇*[i] L1 ≡ K1.ⓑ{I}V1 & K1 ⦻*[R, V1] K2 & R K1 V1 V2.
+#R #L1 #L2 #i #HL12 #I #K2 #V2 #HLK2 elim (lfxs_dropable_dx … HL12 … HLK2 … (#0)) -HLK2 -HL12 //
+#Y #HLK1 #HY elim (lfxs_inv_zero_pair_dx … HY) -HY
+#K1 #V1 #HK12 #HV12 #H destruct /2 width=5 by ex3_2_intro/
+qed-.
index ece33413ebc621f7814bccef71aed79b2c56fe3d..5360aeddc29fd9c35a25b545716b0056d119ff98 100644 (file)
@@ -17,7 +17,7 @@ include "basic_2/static/lfxs.ma".
 
 (* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****)
 
-(* Main properties ******************************************************)
+(* Main properties **********************************************************)
 
 theorem lfxs_bind: ∀R,p,I,L1,L2,V1,V2,T.
                    L1 ⦻*[R, V1] L2 → L1.ⓑ{I}V1 ⦻*[R, T] L2.ⓑ{I}V2 →
index 5b07ac493c3c876ff71c59e3da8cb4409702a695..397134959ef706dad5b5ff813a93ad576fdfa63c 100644 (file)
@@ -144,7 +144,7 @@ table {
         ]
 *)
         [ { "t-bound context-sensitive rt-transition" * } {
-             [ "lfpr ( ⦃?,?⦄ ⊢ ➡[?,?] ? )" "lfpr_length" + "lfpr_drops" + "lfpr_fqup" + "lfpr_lfpx" * ]
+             [ "lfpr ( ⦃?,?⦄ ⊢ ➡[?,?] ? )" "lfpr_length" + "lfpr_drops" + "lfpr_fqup" + "lfpr_lfpx" + "lfpr_lfpr" * ]
              [ "cpr ( ⦃?,?⦄ ⊢ ? ➡[?] ? )" "cpr_drops" (* + "cpr_llpx_sn" + "cpr_cir" *) * ]
              [ "cpm ( ⦃?,?⦄ ⊢ ? ➡[?,?] ? )" "cpm_simple" + "cpm_drops" + "cpm_lsubr" + "cpm_cpx" * ]
           }