]> matita.cs.unibo.it Git - helm.git/commitdiff
- advances on csx
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Thu, 2 Mar 2017 16:00:48 +0000 (16:00 +0000)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Thu, 2 Mar 2017 16:00:48 +0000 (16:00 +0000)
- corrections in the web page

matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tdeq.ma
matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_main.ma [deleted file]
matita/matita/contribs/lambdadelta/basic_2/rt_computation/partial.txt
matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_lfdeq.ma [deleted file]
matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/basic_2/rt_transition/partial.txt
matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq.ma
matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq_tdeq.ma
matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl

index 0271355147c99dbe80f9edd5baae151608a9aaf5..c3e2e20658368c7d0185d4dbbd60f891b2b7b3dd 100644 (file)
 include "basic_2/syntax/tdeq_tdeq.ma".
 include "basic_2/rt_computation/cpxs.ma".
 include "basic_2/rt_transition/cpx_lfdeq.ma".
-include "basic_2/static/lfdeq_fqup.ma".
 include "basic_2/rt_transition/lfpx_fqup.ma".
 
 (* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS ************)
 
-axiom tdeq_dec: ∀h,o,T1,T2. Decidable (tdeq h o T1 T2).
-
-axiom tdeq_canc_sn: ∀h,o. left_cancellable … (tdeq h o).
-
-lemma tdeq_cpx_trans: ∀h,o,U1,T1. U1 ≡[h, o] T1 → ∀G,L,T2. ⦃G, L⦄ ⊢ T1 ⬈[h] T2 → 
-                      ∃∃U2.  ⦃G, L⦄ ⊢ U1 ⬈[h] U2 & U2 ≡[h, o] T2.
-#h #o #U1 #T1 #HUT1 #G #L #T2 #HT12
-elim (cpx_tdeq_conf_lexs … o … HT12 … U1 … L … L) /3 width=3 by tdeq_sym, ex2_intro/
-qed-.
-
 lemma tdeq_cpxs_trans: ∀h,o,U1,T1. U1 ≡[h, o] T1 → ∀G,L,T2. ⦃G, L⦄ ⊢ T1 ⬈*[h] T2 → 
                        ∃∃U2.  ⦃G, L⦄ ⊢ U1 ⬈*[h] U2 & U2 ≡[h, o] T2.
 #h #o #U1 #T1 #HUT1 #G #L #T2 #HT12 @(cpxs_ind … HT12) -T2 /2 width=3 by ex2_intro/
diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx.ma
new file mode 100644 (file)
index 0000000..e1df7c0
--- /dev/null
@@ -0,0 +1,51 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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/syntax/tdeq_tdeq.ma".
+include "basic_2/rt_transition/lfpx_lfdeq.ma".
+include "basic_2/rt_computation/csx.ma".
+
+(* STRONGLY NORMALIZING TERMS FOR UNCOUNTED PARALLEL RT-TRANSITION **********)
+
+(* Advanced properties ******************************************************)
+
+lemma csx_tdeq_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T1⦄ →
+                      ∀T2. T1 ≡[h, o] T2 → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T2⦄.
+#h #o #G #L #T1 #H @(csx_ind … H) -T1 #T #_ #IH #T2 #HT2
+@csx_intro #T1 #HT21 #HnT21 elim (tdeq_cpx_trans … HT2 … HT21) -HT21
+/4 width=5 by tdeq_repl/
+qed-.
+
+lemma csx_cpx_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T1⦄ →
+                     ∀T2. ⦃G, L⦄ ⊢ T1 ⬈[h] T2 → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T2⦄.
+#h #o #G #L #T1 #H @(csx_ind … H) -T1 #T1 #HT1 #IHT1 #T2 #HLT12
+elim (tdeq_dec h o T1 T2) /3 width=4 by csx_tdeq_trans/
+qed-.
+
+(* Basic_1: was just: sn3_cast *)
+lemma csx_cast: ∀h,o,G,L,W. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃W⦄ →
+                ∀T. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃ⓝW.T⦄.
+#h #o #G #L #W #HW @(csx_ind … HW) -W
+#W #HW #IHW #T #HT @(csx_ind … HT) -T
+#T #HT #IHT @csx_intro
+#X #H1 #H2 elim (cpx_inv_cast1 … H1) -H1
+[ * #W0 #T0 #HLW0 #HLT0 #H destruct
+  elim (tdeq_false_inv_pair … H2) -H2
+  [ -W -T #H elim H -H //
+  | -HW -IHT /3 width=3 by csx_cpx_trans/
+  | -HW -HT -IHW /4 width=3 by csx_cpx_trans, cpx_pair_sn/
+  ]
+|*: /3 width=3 by csx_cpx_trans/
+]
+qed.
diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_main.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_main.ma
deleted file mode 100644 (file)
index d3556ce..0000000
+++ /dev/null
@@ -1,34 +0,0 @@
-(*
-lemma csx_tdeq_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T1⦄ →
-                      ∀T2. T1 ≡[h, o] T2 → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T2⦄.
-#h #o #G #L #T1 #H @(csx_ind … H) -H #T #HT #IH #T2 #HT2
-@csx_intro #T0 #HT20 #HnT20      
-
-lemma csx_tdeq_trans: ∀h,o,T1,T2. T1 ≡[h, o] T2 →  
-                      ∀G,L. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T1⦄ → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T2⦄.
-#h #o #T1 #T2 #H elim H -T1 -T2 //
-[ #s1 #s2 #d #Hs1 #Hs2 #G #L #H
-| #I #V1 #V2 #T1 #T2 #_ #_ #IHV #IHT #G #L #H   
-
-lemma csx_cpx_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T1⦄ →
-                     ∀T2. ⦃G, L⦄ ⊢ T1 ⬈[h] T2 → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T2⦄.
-#h #o #G #L #T1 #H @(csx_ind … H) -T1 #T1 #HT1 #IHT1 #T2 #HLT12
-elim (tdeq_dec h o T1 T2) #HT12 /3 width=4 by/ -IHT1 -HLT12
-qed-.
-
-(* Basic_1: was just: sn3_cast *)
-lemma csx_cast: ∀h,o,G,L,W. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃W →
-                ∀T. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃ⓝW.T.
-#h #o #G #L #W #HW @(csx_ind … HW) -W #W #HW #IHW #T #HT @(csx_ind … HT) -T #T #HT #IHT
-@csx_intro #X #H1 #H2
-elim (cpx_inv_cast1 … H1) -H1
-[ * #W0 #T0 #HLW0 #HLT0 #H destruct
-  elim (eq_false_inv_tpair_sn … H2) -H2
-  [ /3 width=3 by csx_cpx_trans/
-  | -HLW0 * #H destruct /3 width=1 by/
-  ]
-|2,3: /3 width=3 by csx_cpx_trans/
-]
-qed.
-
-*)
index fb8f2ebf95e10cba2124a839e8106aee5d0d4c9a..deef34eacf58542a43629ae4d79a4784892dea24 100644 (file)
@@ -1,2 +1,2 @@
 cpxs.ma
-csx.ma csx_cnx.ma
+csx.ma csx_cnx.ma csx_csx.ma
diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_lfdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_lfdeq.ma
deleted file mode 100644 (file)
index 1a4d722..0000000
+++ /dev/null
@@ -1,114 +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/relocation/lifts_tdeq.ma".
-include "basic_2/static/lfdeq.ma".
-include "basic_2/rt_transition/lfpx.ma".
-
-(* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS *************)
-
-(* Properties with degree-based equivalence for terms ***********************)
-
-(* Basic_2A1: was just: cpx_lleq_conf *)
-lemma cpx_tdeq_conf_lexs: ∀h,o,G. R_confluent2_lfxs … (cpx h G) (cdeq h o) (cpx h G) (cdeq h o).
-#h #o #G #L0 #T0 #T1 #H @(cpx_ind … H) -G -L0 -T0 -T1 /2 width=3 by ex2_intro/
-[ #G #L0 #s0 #X0 #H0 #L1 #HL01 #L2 #HL02
-  elim (tdeq_inv_sort1 … H0) -H0 #s1 #d1 #Hs0 #Hs1 #H destruct
-  /4 width=3 by tdeq_sort, deg_next, ex2_intro/
-| #I #G #K0 #V0 #V1 #W1 #_ #IH #HVW1 #T2 #H0 #L1 #H1 #L2 #H2
-  >(tdeq_inv_lref1 … H0) -H0
-  elim (lfpx_inv_zero_pair_sn … H1) -H1 #K1 #X1 #HK01 #HX1 #H destruct
-  elim (lfdeq_inv_zero_pair_sn … H2) -H2 #K2 #X2 #HK02 #HX2 #H destruct
-  elim (IH X2 … HK01 … HK02) // -K0 -V0 #V #HV1 #HV2
-  elim (tdeq_lifts … HV1 … HVW1) -V1 /3 width=5 by cpx_delta, ex2_intro/
-| #I #G #K0 #V0 #V1 #W1 #i #_ #IH #HVW1 #T2 #H0 #L1 #H1 #L2 #H2
-  >(tdeq_inv_lref1 … H0) -H0
-  elim (lfpx_inv_lref_pair_sn … H1) -H1 #K1 #X1 #HK01 #H destruct
-  elim (lfdeq_inv_lref_pair_sn … H2) -H2 #K2 #X2 #HK02 #H destruct
-  elim (IH … HK01 … HK02) [|*: //] -K0 -V0 #V #HV1 #HV2
-  elim (tdeq_lifts … HV1 … HVW1) -V1 /3 width=5 by cpx_lref, ex2_intro/
-| #p #I #G #L0 #V0 #V1 #T0 #T1 #_ #_ #IHV #IHT #X0 #H0 #L1 #H1 #L2 #H2
-  elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #HV02 #HT02 #H destruct
-  elim (lfpx_inv_bind … H1) -H1 #HL01 #H1
-  elim (lfdeq_inv_bind … H2) -H2 #HL02 #H2
-  lapply (lfdeq_pair_repl_dx … H2 … HV02) -H2 #H2
-  elim (IHV … HV02 … HL01 … HL02) -IHV -HV02 -HL01 -HL02
-  elim (IHT … HT02 … H1 … H2) -L0 -T0
-  /3 width=5 by cpx_bind, tdeq_pair, ex2_intro/
-| #I #G #L0 #V0 #V1 #T0 #T1 #_ #_ #IHV #IHT #X0 #H0 #L1 #H1 #L2 #H2
-  elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #HV02 #HT02 #H destruct
-  elim (lfpx_inv_flat … H1) -H1 #HL01 #H1
-  elim (lfdeq_inv_flat … H2) -H2 #HL02 #H2
-  elim (IHV … HV02 … HL01 … HL02) -IHV -HV02 -HL01 -HL02
-  elim (IHT … HT02 … H1 … H2) -L0 -V0 -T0
-  /3 width=5 by cpx_flat, tdeq_pair, ex2_intro/
-| #G #L0 #V0 #T0 #T1 #U1 #_ #IH #HUT1 #X0 #H0 #L1 #H1 #L2 #H2
-  elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #HV02 #HT02 #H destruct
-  elim (lfpx_inv_bind … H1) -H1 #HL01 #H1
-  elim (lfdeq_inv_bind … H2) -H2 #HL02 #H2
-  lapply (lfdeq_pair_repl_dx … H2 … HV02) -H2 -HV02 #H2
-  elim (IH … HT02 … H1 … H2) -L0 -T0 #T #HT1
-  elim (tdeq_inv_lifts … HT1 … HUT1) -T1
-  /3 width=5 by cpx_zeta, ex2_intro/
-| #G #L0 #V0 #T0 #T1 #_ #IH #X0 #H0 #L1 #H1 #L2 #H2
-  elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #_ #HT02 #H destruct
-  elim (lfpx_inv_flat … H1) -H1 #HL01 #H1
-  elim (lfdeq_inv_flat … H2) -H2 #HL02 #H2
-  elim (IH … HT02 … H1 … H2) -L0 -V0 -T0
-  /3 width=3 by cpx_eps, ex2_intro/
-| #G #L0 #V0 #T0 #T1 #_ #IH #X0 #H0 #L1 #H1 #L2 #H2
-  elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #HV02 #_ #H destruct
-  elim (lfpx_inv_flat … H1) -H1 #HL01 #H1
-  elim (lfdeq_inv_flat … H2) -H2 #HL02 #H2
-  elim (IH … HV02 … HL01 … HL02) -L0 -V0 -T1
-  /3 width=3 by cpx_ee, ex2_intro/
-| #p #G #L0 #V0 #V1 #W0 #W1 #T0 #T1 #_ #_ #_ #IHV #IHW #IHT #X0 #H0 #L1 #H1 #L2 #H2
-  elim (tdeq_inv_pair1 … H0) -H0 #V2 #X #HV02 #H0 #H destruct
-  elim (tdeq_inv_pair1 … H0) -H0 #W2 #T2 #HW02 #HT02 #H destruct
-  elim (lfpx_inv_flat … H1) -H1 #H1LV0 #H1
-  elim (lfpx_inv_bind … H1) -H1 #H1LW0 #H1LT0
-  elim (lfdeq_inv_flat … H2) -H2 #H2LV0 #H2
-  elim (lfdeq_inv_bind … H2) -H2 #H2LW0 #H2LT0
-  lapply (lfdeq_pair_repl_dx … H2LT0 … HW02) -H2LT0 #H2LT0
-  elim (IHV … HV02 … H1LV0 … H2LV0) -IHV -HV02 -H1LV0 -H2LV0
-  elim (IHW … HW02 … H1LW0 … H2LW0) -IHW -HW02 -H1LW0 -H2LW0
-  elim (IHT … HT02 … H1LT0 … H2LT0) -L0 -V0 -T0
-  /4 width=7 by cpx_beta, tdeq_pair, ex2_intro/ (* note: 2 tdeq_pair *)
-| #p #G #L0 #V0 #V1 #U1 #W0 #W1 #T0 #T1 #_ #_ #_ #IHV #IHW #IHT #HVU1 #X0 #H0 #L1 #H1 #L2 #H2
-  elim (tdeq_inv_pair1 … H0) -H0 #V2 #X #HV02 #H0 #H destruct
-  elim (tdeq_inv_pair1 … H0) -H0 #W2 #T2 #HW02 #HT02 #H destruct
-  elim (lfpx_inv_flat … H1) -H1 #H1LV0 #H1
-  elim (lfpx_inv_bind … H1) -H1 #H1LW0 #H1LT0
-  elim (lfdeq_inv_flat … H2) -H2 #H2LV0 #H2
-  elim (lfdeq_inv_bind … H2) -H2 #H2LW0 #H2LT0
-  lapply (lfdeq_pair_repl_dx … H2LT0 … HW02) -H2LT0 #H2LT0
-  elim (IHV … HV02 … H1LV0 … H2LV0) -IHV -HV02 -H1LV0 -H2LV0 #V #HV1
-  elim (IHW … HW02 … H1LW0 … H2LW0) -IHW -HW02 -H1LW0 -H2LW0
-  elim (IHT … HT02 … H1LT0 … H2LT0) -L0 -V0 -T0
-  elim (tdeq_lifts … HV1 … HVU1) -V1
-  /4 width=9 by cpx_theta, tdeq_pair, ex2_intro/ (* note: 2 tdeq_pair *)
-]
-qed-.
-(*
-lemma cpx_lleq_conf: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ⬈[h, o] T2 →
-                     ∀L1. L2 ≡[T1, 0] L1 → ⦃G, L1⦄ ⊢ T1 ⬈[h, o] T2.
-/3 width=3 by lleq_cpx_trans, lleq_sym/ qed-.
-
-lemma cpx_lleq_conf_sn: ∀h,o,G. s_r_confluent1 … (cpx h o G) (lleq 0).
-/3 width=6 by cpx_llpx_sn_conf, lift_mono, ex2_intro/ qed-.
-
-lemma cpx_lleq_conf_dx: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ⬈[h, o] T2 →
-                        ∀L1. L1 ≡[T1, 0] L2 → L1 ≡[T2, 0] L2.
-/4 width=6 by cpx_lleq_conf_sn, lleq_sym/ qed-.
-*)
\ No newline at end of file
diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma
new file mode 100644 (file)
index 0000000..d231445
--- /dev/null
@@ -0,0 +1,130 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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/relocation/lifts_tdeq.ma".
+include "basic_2/static/lfdeq_fqup.ma".
+include "basic_2/rt_transition/lfpx.ma".
+
+(* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS *************)
+
+(* Properties with degree-based equivalence for terms ***********************)
+
+(* Basic_2A1: was just: cpx_lleq_conf *)
+lemma cpx_tdeq_conf_lexs: ∀h,o,G. R_confluent2_lfxs … (cpx h G) (cdeq h o) (cpx h G) (cdeq h o).
+#h #o #G #L0 #T0 #T1 #H @(cpx_ind … H) -G -L0 -T0 -T1 /2 width=3 by ex2_intro/
+[ #G #L0 #s0 #X0 #H0 #L1 #HL01 #L2 #HL02
+  elim (tdeq_inv_sort1 … H0) -H0 #s1 #d1 #Hs0 #Hs1 #H destruct
+  /4 width=3 by tdeq_sort, deg_next, ex2_intro/
+| #I #G #K0 #V0 #V1 #W1 #_ #IH #HVW1 #T2 #H0 #L1 #H1 #L2 #H2
+  >(tdeq_inv_lref1 … H0) -H0
+  elim (lfpx_inv_zero_pair_sn … H1) -H1 #K1 #X1 #HK01 #HX1 #H destruct
+  elim (lfdeq_inv_zero_pair_sn … H2) -H2 #K2 #X2 #HK02 #HX2 #H destruct
+  elim (IH X2 … HK01 … HK02) // -K0 -V0 #V #HV1 #HV2
+  elim (tdeq_lifts … HV1 … HVW1) -V1 /3 width=5 by cpx_delta, ex2_intro/
+| #I #G #K0 #V0 #V1 #W1 #i #_ #IH #HVW1 #T2 #H0 #L1 #H1 #L2 #H2
+  >(tdeq_inv_lref1 … H0) -H0
+  elim (lfpx_inv_lref_pair_sn … H1) -H1 #K1 #X1 #HK01 #H destruct
+  elim (lfdeq_inv_lref_pair_sn … H2) -H2 #K2 #X2 #HK02 #H destruct
+  elim (IH … HK01 … HK02) [|*: //] -K0 -V0 #V #HV1 #HV2
+  elim (tdeq_lifts … HV1 … HVW1) -V1 /3 width=5 by cpx_lref, ex2_intro/
+| #p #I #G #L0 #V0 #V1 #T0 #T1 #_ #_ #IHV #IHT #X0 #H0 #L1 #H1 #L2 #H2
+  elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #HV02 #HT02 #H destruct
+  elim (lfpx_inv_bind … H1) -H1 #HL01 #H1
+  elim (lfdeq_inv_bind … H2) -H2 #HL02 #H2
+  lapply (lfdeq_pair_repl_dx … H2 … HV02) -H2 #H2
+  elim (IHV … HV02 … HL01 … HL02) -IHV -HV02 -HL01 -HL02
+  elim (IHT … HT02 … H1 … H2) -L0 -T0
+  /3 width=5 by cpx_bind, tdeq_pair, ex2_intro/
+| #I #G #L0 #V0 #V1 #T0 #T1 #_ #_ #IHV #IHT #X0 #H0 #L1 #H1 #L2 #H2
+  elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #HV02 #HT02 #H destruct
+  elim (lfpx_inv_flat … H1) -H1 #HL01 #H1
+  elim (lfdeq_inv_flat … H2) -H2 #HL02 #H2
+  elim (IHV … HV02 … HL01 … HL02) -IHV -HV02 -HL01 -HL02
+  elim (IHT … HT02 … H1 … H2) -L0 -V0 -T0
+  /3 width=5 by cpx_flat, tdeq_pair, ex2_intro/
+| #G #L0 #V0 #T0 #T1 #U1 #_ #IH #HUT1 #X0 #H0 #L1 #H1 #L2 #H2
+  elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #HV02 #HT02 #H destruct
+  elim (lfpx_inv_bind … H1) -H1 #HL01 #H1
+  elim (lfdeq_inv_bind … H2) -H2 #HL02 #H2
+  lapply (lfdeq_pair_repl_dx … H2 … HV02) -H2 -HV02 #H2
+  elim (IH … HT02 … H1 … H2) -L0 -T0 #T #HT1
+  elim (tdeq_inv_lifts … HT1 … HUT1) -T1
+  /3 width=5 by cpx_zeta, ex2_intro/
+| #G #L0 #V0 #T0 #T1 #_ #IH #X0 #H0 #L1 #H1 #L2 #H2
+  elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #_ #HT02 #H destruct
+  elim (lfpx_inv_flat … H1) -H1 #HL01 #H1
+  elim (lfdeq_inv_flat … H2) -H2 #HL02 #H2
+  elim (IH … HT02 … H1 … H2) -L0 -V0 -T0
+  /3 width=3 by cpx_eps, ex2_intro/
+| #G #L0 #V0 #T0 #T1 #_ #IH #X0 #H0 #L1 #H1 #L2 #H2
+  elim (tdeq_inv_pair1 … H0) -H0 #V2 #T2 #HV02 #_ #H destruct
+  elim (lfpx_inv_flat … H1) -H1 #HL01 #H1
+  elim (lfdeq_inv_flat … H2) -H2 #HL02 #H2
+  elim (IH … HV02 … HL01 … HL02) -L0 -V0 -T1
+  /3 width=3 by cpx_ee, ex2_intro/
+| #p #G #L0 #V0 #V1 #W0 #W1 #T0 #T1 #_ #_ #_ #IHV #IHW #IHT #X0 #H0 #L1 #H1 #L2 #H2
+  elim (tdeq_inv_pair1 … H0) -H0 #V2 #X #HV02 #H0 #H destruct
+  elim (tdeq_inv_pair1 … H0) -H0 #W2 #T2 #HW02 #HT02 #H destruct
+  elim (lfpx_inv_flat … H1) -H1 #H1LV0 #H1
+  elim (lfpx_inv_bind … H1) -H1 #H1LW0 #H1LT0
+  elim (lfdeq_inv_flat … H2) -H2 #H2LV0 #H2
+  elim (lfdeq_inv_bind … H2) -H2 #H2LW0 #H2LT0
+  lapply (lfdeq_pair_repl_dx … H2LT0 … HW02) -H2LT0 #H2LT0
+  elim (IHV … HV02 … H1LV0 … H2LV0) -IHV -HV02 -H1LV0 -H2LV0
+  elim (IHW … HW02 … H1LW0 … H2LW0) -IHW -HW02 -H1LW0 -H2LW0
+  elim (IHT … HT02 … H1LT0 … H2LT0) -L0 -V0 -T0
+  /4 width=7 by cpx_beta, tdeq_pair, ex2_intro/ (* note: 2 tdeq_pair *)
+| #p #G #L0 #V0 #V1 #U1 #W0 #W1 #T0 #T1 #_ #_ #_ #IHV #IHW #IHT #HVU1 #X0 #H0 #L1 #H1 #L2 #H2
+  elim (tdeq_inv_pair1 … H0) -H0 #V2 #X #HV02 #H0 #H destruct
+  elim (tdeq_inv_pair1 … H0) -H0 #W2 #T2 #HW02 #HT02 #H destruct
+  elim (lfpx_inv_flat … H1) -H1 #H1LV0 #H1
+  elim (lfpx_inv_bind … H1) -H1 #H1LW0 #H1LT0
+  elim (lfdeq_inv_flat … H2) -H2 #H2LV0 #H2
+  elim (lfdeq_inv_bind … H2) -H2 #H2LW0 #H2LT0
+  lapply (lfdeq_pair_repl_dx … H2LT0 … HW02) -H2LT0 #H2LT0
+  elim (IHV … HV02 … H1LV0 … H2LV0) -IHV -HV02 -H1LV0 -H2LV0 #V #HV1
+  elim (IHW … HW02 … H1LW0 … H2LW0) -IHW -HW02 -H1LW0 -H2LW0
+  elim (IHT … HT02 … H1LT0 … H2LT0) -L0 -V0 -T0
+  elim (tdeq_lifts … HV1 … HVU1) -V1
+  /4 width=9 by cpx_theta, tdeq_pair, ex2_intro/ (* note: 2 tdeq_pair *)
+]
+qed-.
+
+lemma cpx_tdeq_conf: ∀h,o,G,L,T0,T1. ⦃G, L⦄ ⊢ T0 ⬈[h] T1 →
+                     ∀T2. T0 ≡[h, o] T2 →
+                     ∃∃T. T1 ≡[h, o] T & ⦃G, L⦄ ⊢ T2 ⬈[h] T.
+#h #o #G #L #T0 #T1 #HT01 #T2 #HT02
+elim (cpx_tdeq_conf_lexs … HT01 … HT02 L … L) -HT01 -HT02
+/2 width=3 by lfxs_refl, ex2_intro/
+qed-.
+
+lemma tdeq_cpx_trans: ∀h,o,G,L,T2,T0. T2 ≡[h, o] T0 →
+                      ∀T1. ⦃G, L⦄ ⊢ T0 ⬈[h] T1 → 
+                      ∃∃T. ⦃G, L⦄ ⊢ T2 ⬈[h] T & T ≡[h, o] T1.
+#h #o #G #L #T2 #T0 #HT20 #T1 #HT01
+elim (cpx_tdeq_conf … HT01 T2) -HT01 /3 width=3 by tdeq_sym, ex2_intro/
+qed-.
+
+(*
+lemma cpx_lleq_conf: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ⬈[h, o] T2 →
+                     ∀L1. L2 ≡[T1, 0] L1 → ⦃G, L1⦄ ⊢ T1 ⬈[h, o] T2.
+/3 width=3 by lleq_cpx_trans, lleq_sym/ qed-.
+
+lemma cpx_lleq_conf_sn: ∀h,o,G. s_r_confluent1 … (cpx h o G) (lleq 0).
+/3 width=6 by cpx_llpx_sn_conf, lift_mono, ex2_intro/ qed-.
+
+lemma cpx_lleq_conf_dx: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ⬈[h, o] T2 →
+                        ∀L1. L1 ≡[T1, 0] L2 → L1 ≡[T2, 0] L2.
+/4 width=6 by cpx_lleq_conf_sn, lleq_sym/ qed-.
+*)
\ No newline at end of file
index 9ed6e41045337ef0e0d5001038ab737b27b8595a..1adc75c386bad4d5c0b85cc00216d877eb232c15 100644 (file)
@@ -1,6 +1,6 @@
 cpg.ma cpg_simple.ma cpg_drops.ma cpg_lsubr.ma
-cpx.ma cpx_simple.ma cpx_drops.ma cpx_fqus.ma cpx_lsubr.ma cpx_lfxs.ma cpx_lfdeq.ma
-lfpx.ma lfpx_length.ma lfpx_drops.ma lfpx_fqup.ma lfpx_frees.ma lfpx_aaa.ma
+cpx.ma cpx_simple.ma cpx_drops.ma cpx_fqus.ma cpx_lsubr.ma cpx_lfxs.ma
+lfpx.ma lfpx_length.ma lfpx_drops.ma lfpx_fqup.ma lfpx_frees.ma lfpx_lfdeq.ma lfpx_aaa.ma
 cnx.ma cnx_simple.ma cnx_drops.ma
 cpm.ma cpm_simple.ma cpm_drops.ma cpm_lsubr.ma cpm_lfxs.ma cpm_cpx.ma
 cpr.ma cpr_drops.ma
index 4e5e41b4db741f9e53392f332b32e630c2df1304..4b077059d51a06a5f7ce9b56f38616dff747b2db 100644 (file)
@@ -168,3 +168,16 @@ lemma tdeq_dec: ∀h,o,T1,T2. Decidable (tdeq h o T1 T2).
   elim (tdeq_inv_pair … H) -H /2 width=1 by/
 ]
 qed-.
+
+(* Negated inversion lemmas *************************************************)
+
+lemma tdeq_false_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 #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/
+elim (tdeq_dec h o T1 T2) /4 width=1 by tdeq_pair, or3_intro2/
+qed-.
index 39b06145cd3706f2c5438b48fe4001d1bc1c7ecd..9f47fd88f5fbeb8370b45f1f88bfef0b184c5f7d 100644 (file)
@@ -28,3 +28,13 @@ theorem tdeq_trans: ∀h,o. Transitive … (tdeq h o).
   elim (tdeq_inv_pair1 … H) -H /3 width=1 by tdeq_pair/
 ]
 qed-.
+
+theorem tdeq_canc_sn: ∀h,o. left_cancellable … (tdeq h o).
+/3 width=3 by tdeq_trans, tdeq_sym/ qed-.
+
+theorem tdeq_canc_dx: ∀h,o. right_cancellable … (tdeq h o).
+/3 width=3 by tdeq_trans, tdeq_sym/ qed-.
+
+theorem tdeq_repl: ∀h,o,T1,T2. T1 ≡[h, o] T2 →
+                   ∀U1. T1 ≡[h, o] U1 → ∀U2. T2 ≡[h, o] U2 → U1 ≡[h, o] U2.
+/3 width=3 by tdeq_canc_sn, tdeq_trans/ qed-.
index 4a5a696aca5db847df63cdf3e6ca39d468218e0c..267559ff57a6ce73a0ae3cdb8c81e53ea2d06298 100644 (file)
@@ -121,6 +121,7 @@ table {
              [ "lpxs ( ⦃?,?⦄ ⊢ ➡*[?,?] ? )" "lpxs_drop" + "lpxs_lleq" + "lpxs_aaa" + "lpxs_cpxs" + "lpxs_lpxs" * ]
              [ "cpxs_tsts" + "cpxs_tsts_vector" + "cpxs_lreq" + "cpxs_lift" + "cpxs_lleq" + "cpxs_aaa" + "cpxs_cpxs" * ]
 *)
+             [ "csx ( ⦃?,?⦄ ⊢ ⬈[?,?] 𝐒⦃?⦄ )" "csx_cnx" + "csx_csx" * ]
              [ "cpxs ( ⦃?,?⦄ ⊢ ? ⬈*[?] ? )" * ] 
           }
         ]
@@ -128,7 +129,7 @@ table {
    ]
    class "water"
    [ { "rt-transition" * } {
-        [ { "parallel qrst-rtransition" * } {
+        [ { "parallel qrst-transition" * } {
 (*             [ "fpbq ( ⦃?,?,?⦄ ≽[?,?] ⦃?,?,?⦄ )" "fpbq_alt ( ⦃?,?,?⦄ ≽≽[?,?] ⦃?,?,?⦄ )"  + "fpbq_aaa" * ] *)
              [ "fpb ( ⦃?,?,?⦄ ≻[?,?] ⦃?,?,?⦄ )" (* "fpb_lleq" + "fpb_fleq" *) * ]
           }
@@ -148,8 +149,8 @@ table {
         ]
         [ { "uncounted context-sensitive rt-transition" * } {
              [ "cnx ( ⦃?,?⦄ ⊢ ⬈[?,?] 𝐍⦃?⦄ )" "cnx_simple" + "cnx_drops" * ]
-             [ "lfpx ( ⦃?,?⦄ ⊢ ⬈[?,?] ? )" "lfpx_length" + "lfpx_drops" + "lfpx_fqup" + "lfpx_frees" + "lfpx_aaa" * ]
-             [ "cpx ( ⦃?,?⦄ ⊢ ? ⬈[?] ? )" "cpx_simple" + "cpx_drops" + "cpx_fqus" + "cpx_lsubr" + "cpx_lfxs" + "cpx_lfdeq" * ]
+             [ "lfpx ( ⦃?,?⦄ ⊢ ⬈[?,?] ? )" "lfpx_length" + "lfpx_drops" + "lfpx_fqup" + "lfpx_frees" + "lfpx_lfdeq" + "lfpx_aaa" * ]
+             [ "cpx ( ⦃?,?⦄ ⊢ ? ⬈[?] ? )" "cpx_simple" + "cpx_drops" + "cpx_fqus" + "cpx_lsubr" + "cpx_lfxs" * ]
           }
         ]
         [ { "counted context-sensitive rt-transition" * } {
@@ -161,7 +162,7 @@ table {
    class "green"
    [ { "static typing" * } {
         [ { "generic reducibility" * } {
-             [ "lsubc ( ? ⊢ ? ⫃[?] ? )" "lsubc_drop" + "lsubc_drops" + "lsubc_lsubr" + "lsubc_lsuba" * ]        
+             [ "lsubc ( ? ⊢ ? ⫃[?] ? )" "lsubc_drops" + "lsubc_lsubr" + "lsubc_lsuba" * ]        
              [ "gcp_cr ( ⦃?,?,?⦄ ϵ[?] 〚?〛 )" "gcp_aaa" * ]
              [ "gcp" *] 
           }
@@ -182,7 +183,7 @@ table {
         ]
         [ { "context-sensitive free variables" * } {
              [ "lsubf ( ⦃?,?⦄ ⫃𝐅* ⦃?,?⦄ )" "lsubf_frees" * ]
-             [ "frees ( ? ⊢ 𝐅*⦃?⦄ ≡ ? )" "frees_weight" + "frees_drops" + "frees_fqup" + "frees_frees" * ]
+             [ "frees ( ? ⊢ 𝐅*⦃?⦄ ≡ ? )" "frees_drops" + "frees_fqup" + "frees_frees" * ]
           }
         ]
         [ { "restricted ref. for local env." * } {