]> matita.cs.unibo.it Git - helm.git/commitdiff
- strongly normalizing terms form a candidate of reducibility
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Sat, 1 Apr 2017 13:03:45 +0000 (13:03 +0000)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Sat, 1 Apr 2017 13:03:45 +0000 (13:03 +0000)
- one application of auto is unnecessarily slow,
  a longly awaited improvement of auto is needed to solve the issue

matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_etc.ma [deleted file]
matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_fqus.ma
matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_lfpx.ma
matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx_vector.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_gcp.ma
matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_gcr.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_theq_vector.ma [deleted file]
matita/matita/contribs/lambdadelta/basic_2/rt_computation/partial.txt
matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl

diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_etc.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_etc.ma
deleted file mode 100644 (file)
index 0340000..0000000
+++ /dev/null
@@ -1,81 +0,0 @@
-
-include "basic_2/reduction/lpx_drop.ma".
-include "basic_2/computation/cpxs_lift.ma".
-include "basic_2/rt_computation/cpxs_cpxs.ma".
-
-(* Properties on sn extended parallel reduction for local environments ******)
-
-lemma cpx_bind2: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ⬈[h, o] V2 →
-                 ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ⬈[h, o] T2 →
-                 ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ⬈*[h, o] ⓑ{a,I}V2.T2.
-/4 width=5 by lpx_cpx_trans, cpxs_bind_dx, lpx_pair/ qed.
-
-(* Advanced properties ******************************************************)
-
-lemma cpxs_bind2_dx: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ⬈[h, o] V2 →
-                     ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ⬈*[h, o] T2 →
-                     ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ⬈*[h, o] ⓑ{a,I}V2.T2.
-/4 width=5 by lpx_cpxs_trans, cpxs_bind_dx, lpx_pair/ qed.
-
-(* Properties on supclosure *************************************************)
-
-lemma fqu_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
-                          ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈*[h, o] U2 → (T2 = U2 → ⊥) →
-                          ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
-[ #I #G #L #V1 #V2 #HV12 #_ elim (lift_total V2 0 1)
-  #U2 #HVU2 @(ex3_intro … U2)
-  [1,3: /3 width=7 by fqu_drop, cpxs_delta, drop_pair, drop_drop/
-  | #H destruct 
-    lapply (lift_inv_lref2_be … HVU2 ? ?) -HVU2 //
-  ]
-| #I #G #L #V1 #T #V2 #HV12 #H @(ex3_intro … (②{I}V2.T))
-  [1,3: /2 width=4 by fqu_pair_sn, cpxs_pair_sn/
-  | #H0 destruct /2 width=1 by/
-  ]
-| #a #I #G #L #V #T1 #T2 #HT12 #H @(ex3_intro … (ⓑ{a,I}V.T2))
-  [1,3: /2 width=4 by fqu_bind_dx, cpxs_bind/
-  | #H0 destruct /2 width=1 by/
-  ]
-| #I #G #L #V #T1 #T2 #HT12 #H @(ex3_intro … (ⓕ{I}V.T2))
-  [1,3: /2 width=4 by fqu_flat_dx, cpxs_flat/
-  | #H0 destruct /2 width=1 by/
-  ]
-| #G #L #K #T1 #U1 #k #HLK #HTU1 #T2 #HT12 #H elim (lift_total T2 0 (k+1))
-  #U2 #HTU2 @(ex3_intro … U2)
-  [1,3: /2 width=10 by cpxs_lift, fqu_drop/
-  | #H0 destruct /3 width=5 by lift_inj/
-]
-qed-.
-
-lemma fquq_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
-                           ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈*[h, o] U2 → (T2 = U2 → ⊥) →
-                           ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fquq_inv_gen … H12) -H12
-[ #H12 elim (fqu_cpxs_trans_neq … H12 … HTU2 H) -T2
-  /3 width=4 by fqu_fquq, ex3_intro/
-| * #HG #HL #HT destruct /3 width=4 by ex3_intro/
-]
-qed-.
-
-lemma fqup_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
-                           ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈*[h, o] U2 → (T2 = U2 → ⊥) →
-                           ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind_dx … H) -G1 -L1 -T1
-[ #G1 #L1 #T1 #H12 #U2 #HTU2 #H elim (fqu_cpxs_trans_neq … H12 … HTU2 H) -T2
-  /3 width=4 by fqu_fqup, ex3_intro/
-| #G #G1 #L #L1 #T #T1 #H1 #_ #IH12 #U2 #HTU2 #H elim (IH12 … HTU2 H) -T2
-  #U1 #HTU1 #H #H12 elim (fqu_cpxs_trans_neq … H1 … HTU1 H) -T1
-  /3 width=8 by fqup_strap2, ex3_intro/
-]
-qed-.
-
-lemma fqus_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
-                           ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈*[h, o] U2 → (T2 = U2 → ⊥) →
-                           ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fqus_inv_gen … H12) -H12
-[ #H12 elim (fqup_cpxs_trans_neq … H12 … HTU2 H) -T2
-  /3 width=4 by fqup_fqus, ex3_intro/
-| * #HG #HL #HT destruct /3 width=4 by ex3_intro/
-]
-qed-.
index 4bacb47f6d7457db6462dc2a3ffcaf5e39bd11f5..82e21018e18e735dfd149b3a2bc3ec9443ee9a0f 100644 (file)
@@ -68,3 +68,66 @@ lemma fqus_lstas_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L
                         ∀d2. ⦃G2, L2⦄ ⊢ T2 ▪[h, o] d2 → d1 ≤ d2 →
                         ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
 /3 width=6 by fqus_cpxs_trans, lstas_cpxs/ qed-.
+
+(* Properties on supclosure *************************************************)
+
+lemma fqu_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+                          ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈*[h, o] U2 → (T2 = U2 → ⊥) →
+                          ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+[ #I #G #L #V1 #V2 #HV12 #_ elim (lift_total V2 0 1)
+  #U2 #HVU2 @(ex3_intro … U2)
+  [1,3: /3 width=7 by fqu_drop, cpxs_delta, drop_pair, drop_drop/
+  | #H destruct 
+    lapply (lift_inv_lref2_be … HVU2 ? ?) -HVU2 //
+  ]
+| #I #G #L #V1 #T #V2 #HV12 #H @(ex3_intro … (②{I}V2.T))
+  [1,3: /2 width=4 by fqu_pair_sn, cpxs_pair_sn/
+  | #H0 destruct /2 width=1 by/
+  ]
+| #a #I #G #L #V #T1 #T2 #HT12 #H @(ex3_intro … (ⓑ{a,I}V.T2))
+  [1,3: /2 width=4 by fqu_bind_dx, cpxs_bind/
+  | #H0 destruct /2 width=1 by/
+  ]
+| #I #G #L #V #T1 #T2 #HT12 #H @(ex3_intro … (ⓕ{I}V.T2))
+  [1,3: /2 width=4 by fqu_flat_dx, cpxs_flat/
+  | #H0 destruct /2 width=1 by/
+  ]
+| #G #L #K #T1 #U1 #k #HLK #HTU1 #T2 #HT12 #H elim (lift_total T2 0 (k+1))
+  #U2 #HTU2 @(ex3_intro … U2)
+  [1,3: /2 width=10 by cpxs_lift, fqu_drop/
+  | #H0 destruct /3 width=5 by lift_inj/
+]
+qed-.
+
+lemma fquq_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+                           ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈*[h, o] U2 → (T2 = U2 → ⊥) →
+                           ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fquq_inv_gen … H12) -H12
+[ #H12 elim (fqu_cpxs_trans_neq … H12 … HTU2 H) -T2
+  /3 width=4 by fqu_fquq, ex3_intro/
+| * #HG #HL #HT destruct /3 width=4 by ex3_intro/
+]
+qed-.
+
+lemma fqup_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+                           ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈*[h, o] U2 → (T2 = U2 → ⊥) →
+                           ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind_dx … H) -G1 -L1 -T1
+[ #G1 #L1 #T1 #H12 #U2 #HTU2 #H elim (fqu_cpxs_trans_neq … H12 … HTU2 H) -T2
+  /3 width=4 by fqu_fqup, ex3_intro/
+| #G #G1 #L #L1 #T #T1 #H1 #_ #IH12 #U2 #HTU2 #H elim (IH12 … HTU2 H) -T2
+  #U1 #HTU1 #H #H12 elim (fqu_cpxs_trans_neq … H1 … HTU1 H) -T1
+  /3 width=8 by fqup_strap2, ex3_intro/
+]
+qed-.
+
+lemma fqus_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+                           ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈*[h, o] U2 → (T2 = U2 → ⊥) →
+                           ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fqus_inv_gen … H12) -H12
+[ #H12 elim (fqup_cpxs_trans_neq … H12 … HTU2 H) -T2
+  /3 width=4 by fqup_fqus, ex3_intro/
+| * #HG #HL #HT destruct /3 width=4 by ex3_intro/
+]
+qed-.
index b54ea1582eaf871673fec61685f90743b0876c5b..8c92f374235b46b64923f2acbd4b3f95537b440e 100644 (file)
@@ -12,6 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
+include "basic_2/rt_transition/lfpx_fqup.ma".
 include "basic_2/rt_transition/lfpx_lfpx.ma".
 include "basic_2/rt_computation/cpxs_drops.ma".
 include "basic_2/rt_computation/cpxs_cpxs.ma".
@@ -53,5 +54,17 @@ qed.
 
 lemma lfpx_cpxs_trans: ∀h,G. s_rs_transitive … (cpx h G) (lfpx h G).
 #h #G #L2 #T1 #T2 #H #L1 #HL12 @(cpxs_ind … H) -T2
-/4 width=7 by lfpx_cpx_trans, cpxs_trans, lfpx_cpxs_conf/
+/4 width=7 by lfpx_cpx_trans, cpxs_trans, lfpx_cpxs_conf/ (**) (* lfpx_fqup slows this down *)
 qed-.
+
+(* Advanced properties ******************************************************)
+
+lemma cpx_bind2: ∀h,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ⬈[h] V2 →
+                 ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ⬈[h] T2 →
+                 ∀p. ⦃G, L⦄ ⊢ ⓑ{p,I}V1.T1 ⬈*[h] ⓑ{p,I}V2.T2.
+/4 width=5 by lfpx_cpx_trans, cpxs_bind_dx, lfpx_pair/ qed.
+
+lemma cpxs_bind2_dx: ∀h,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ⬈[h] V2 →
+                     ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ⬈*[h] T2 →
+                     ∀p. ⦃G, L⦄ ⊢ ⓑ{p,I}V1.T1 ⬈*[h] ⓑ{p,I}V2.T2.
+/4 width=5 by lfpx_cpxs_trans, cpxs_bind_dx, lfpx_pair/ qed.
diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx_vector.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx_vector.ma
new file mode 100644 (file)
index 0000000..fa81a32
--- /dev/null
@@ -0,0 +1,91 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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_computation/cpxs_theq_vector.ma".
+include "basic_2/rt_computation/csx_simple_theq.ma".
+include "basic_2/rt_computation/csx_lsubr.ma".
+include "basic_2/rt_computation/csx_lfpx.ma".
+include "basic_2/rt_computation/csx_vector.ma".
+
+(* STRONGLY NORMALIZING TERM VECTORS FOR UNCOUNTED PARALLEL RT-TRANSITION ***)
+
+(* Advanced properties ************************************* ****************)
+
+(* Basic_1: was just: sn3_appls_beta *)
+lemma csx_applv_beta: ∀h,o,p,G,L,Vs,V,W,T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⒶVs.ⓓ{p}ⓝW.V.T⦄ →
+                      ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⒶVs.ⓐV.ⓛ{p}W.T⦄.
+#h #o #p #G #L #Vs elim Vs -Vs /2 width=1 by csx_appl_beta/
+#V0 #Vs #IHV #V #W #T #H1T
+lapply (csx_fwd_pair_sn … H1T) #HV0
+lapply (csx_fwd_flat_dx … H1T) #H2T
+@csx_appl_simple_theq /2 width=1 by applv_simple, simple_flat/ -IHV -HV0 -H2T
+#X #H #H0
+elim (cpxs_fwd_beta_vector … o … H) -H #H
+[ -H1T elim H0 -H0 //
+| -H0 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
+]
+qed.
+
+lemma csx_applv_delta: ∀h,o,I,G,L,K,V1,i. ⬇*[i] L ≡ K.ⓑ{I}V1 →
+                       ∀V2. ⬆*[⫯i] V1 ≡ V2 →
+                       ∀Vs. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⒶVs.V2⦄ → ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⒶVs.#i⦄.
+#h #o #I #G #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs
+[ /4 width=11 by csx_inv_lifts, csx_lref_drops, drops_isuni_fwd_drop2/
+| #V #Vs #IHV #H1T
+  lapply (csx_fwd_pair_sn … H1T) #HV
+  lapply (csx_fwd_flat_dx … H1T) #H2T
+  @csx_appl_simple_theq /2 width=1 by applv_simple, simple_atom/ -IHV -HV  -H2T
+  #X #H #H0
+  elim (cpxs_fwd_delta_drops_vector … o … HLK … HV12 … H) -HLK -HV12 -H #H
+  [ -H1T elim H0 -H0 //
+  | -H0 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
+  ]
+]
+qed.
+
+(* Basic_1: was just: sn3_appls_abbr *)
+lemma csx_applv_theta: ∀h,o,p,G,L,V1b,V2b. ⬆*[1] V1b ≡ V2b →
+                       ∀V,T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⓓ{p}V.ⒶV2b.T⦄ →
+                       ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⒶV1b.ⓓ{p}V.T⦄.
+#h #o #p #G #L #V1b #V2b * -V1b -V2b /2 width=1 by/
+#V1b #V2b #V1 #V2 #HV12 #H
+generalize in match HV12; -HV12 generalize in match V2; -V2 generalize in match V1; -V1
+elim H -V1b -V2b /2 width=3 by csx_appl_theta/
+#V1b #V2b #V1 #V2 #HV12 #HV12b #IHV12b #W1 #W2 #HW12 #V #T #H
+lapply (csx_appl_theta … H … HW12) -H -HW12 #H
+lapply (csx_fwd_pair_sn … H) #HW1
+lapply (csx_fwd_flat_dx … H) #H1
+@csx_appl_simple_theq /2 width=3 by simple_flat/ -IHV12b -HW1 -H1 #X #H1 #H2
+elim (cpxs_fwd_theta_vector … o … (V2@V2b) … H1) -H1 /2 width=1 by liftsv_cons/ -HV12b -HV12
+[ -H #H elim H2 -H2 //
+| -H2 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
+]
+qed.
+
+(* Basic_1: was just: sn3_appls_cast *)
+lemma csx_applv_cast: ∀h,o,G,L,Vs,U. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⒶVs.U⦄ →
+                      ∀T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⒶVs.T⦄ → ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⒶVs.ⓝU.T⦄.
+#h #o #G #L #Vs elim Vs -Vs /2 width=1 by csx_cast/
+#V #Vs #IHV #U #H1U #T #H1T
+lapply (csx_fwd_pair_sn … H1U) #HV
+lapply (csx_fwd_flat_dx … H1U) #H2U
+lapply (csx_fwd_flat_dx … H1T) #H2T
+@csx_appl_simple_theq /2 width=1 by applv_simple, simple_flat/ -IHV -HV -H2U -H2T
+#X #H #H0
+elim (cpxs_fwd_cast_vector … o … H) -H #H
+[ -H1U -H1T elim H0 -H0 //
+| -H1U -H0 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
+| -H1T -H0 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
+]
+qed.
index a5bfe6cdf008ddfae45a4c7b49d5cd512b51ac0f..5a0426628a6f19bbb7518e5d5d253c9699433de1 100644 (file)
@@ -18,7 +18,7 @@ include "basic_2/rt_computation/csx_drops.ma".
 
 (* STRONGLY NORMALIZING TERMS FOR UNCOUNTED PARALLEL RT-TRANSITION **********)
 
-(* Main properties **********************************************************)
+(* Main properties with generic computation properties **********************)
 
 theorem csx_gcp: ∀h,o. gcp (cpx h) (tdeq h o) (csx h o).
 #h #o @mk_gcp
diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_gcr.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_gcr.ma
new file mode 100644 (file)
index 0000000..f1f1393
--- /dev/null
@@ -0,0 +1,32 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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/static/gcp_cr.ma".
+include "basic_2/rt_computation/csx_cnx_vector.ma".
+include "basic_2/rt_computation/csx_csx_vector.ma".
+
+(* STRONGLY NORMALIZING TERMS FOR UNCOUNTED PARALLEL RT-TRANSITION **********)
+
+(* Main properties with generic candidates of reducibility ******************)
+
+theorem csx_gcr: ∀h,o. gcr (cpx h) (tdeq h o) (csx h o) (csx h o).
+#h #o @mk_gcr //
+[ /3 width=1 by csx_applv_cnx/
+|2,3,6: /2 width=1 by csx_applv_beta, csx_applv_sort, csx_applv_cast/
+| /2 width=7 by csx_applv_delta/
+| #G #L #V1b #V2b #HV12b #a #V #T #H #HV
+  @(csx_applv_theta … HV12b) -HV12b
+  @csx_abbr //
+]
+qed.
diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_theq_vector.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_theq_vector.ma
deleted file mode 100644 (file)
index 55146e3..0000000
+++ /dev/null
@@ -1,107 +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/computation/gcp_cr.ma".
-*)
-include "basic_2/rt_computation/cpxs_theq_vector.ma".
-include "basic_2/rt_computation/csx_vector.ma".
-include "basic_2/rt_computation/csx_theq.ma".
-include "basic_2/rt_computation/csx_lfpx.ma".
-
-(* STRONGLY NORMALIZING TERMS FOR UNCOUNTED PARALLEL RT-TRANSITION **********)
-
-(* Vector form of properties with head equivalence for terms ****************)
-(*
-*)
-(*
-*)
-
-(* Basic_1: was just: sn3_appls_beta *)
-lemma csx_applv_beta: ∀h,o,p,G,L,Vs,V,W,T. ⦃G, L⦄ ⊢ ⬈*[h, o] ⒶVs.ⓓ{p}ⓝW.V.T →
-                      ⦃G, L⦄ ⊢ ⬈*[h, o] ⒶVs. ⓐV.ⓛ{p}W.T.
-#h #o #a #G #L #Vs elim Vs -Vs /2 width=1 by csx_appl_beta/
-#V0 #Vs #IHV #V #W #T #H1T
-lapply (csx_fwd_pair_sn … H1T) #HV0
-lapply (csx_fwd_flat_dx … H1T) #H2T
-@csx_appl_simple_theq /2 width=1 by applv_simple, simple_flat/ -IHV -HV0 -H2T
-#X #H #H0
-elim (cpxs_fwd_beta_vector … H) -H #H
-[ -H1T elim H0 -H0 //
-| -H0 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
-]
-qed.
-
-lemma csx_applv_delta: ∀h,o,I,G,L,K,V1,i. ⬇[i] L ≡ K.ⓑ{I}V1 →
-                       ∀V2. ⬆[0, i + 1] V1 ≡ V2 →
-                       ∀Vs. ⦃G, L⦄ ⊢ ⬈*[h, o] (ⒶVs.V2) → ⦃G, L⦄ ⊢ ⬈*[h, o] (ⒶVs.#i).
-#h #o #I #G #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs
-[ /4 width=12 by csx_inv_lift, csx_lref_bind, drop_fwd_drop2/
-| #V #Vs #IHV #H1T
-  lapply (csx_fwd_pair_sn … H1T) #HV
-  lapply (csx_fwd_flat_dx … H1T) #H2T
-  @csx_appl_simple_theq /2 width=1 by applv_simple, simple_atom/ -IHV -HV  -H2T
-  #X #H #H0
-  elim (cpxs_fwd_delta_vector … HLK … HV12 … H) -HLK -HV12 -H #H
-  [ -H1T elim H0 -H0 //
-  | -H0 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
-  ]
-]
-qed.
-
-(* Basic_1: was just: sn3_appls_abbr *)
-lemma csx_applv_theta: ∀h,o,a,G,L,V1b,V2b. ⬆[0, 1] V1b ≡ V2b →
-                       ∀V,T. ⦃G, L⦄ ⊢ ⬈*[h, o] ⓓ{a}V.ⒶV2b.T →
-                       ⦃G, L⦄ ⊢ ⬈*[h, o] ⒶV1b.ⓓ{a}V.T.
-#h #o #a #G #L #V1b #V2b * -V1b -V2b /2 width=1 by/
-#V1b #V2b #V1 #V2 #HV12 #H
-generalize in match HV12; -HV12 generalize in match V2; -V2 generalize in match V1; -V1
-elim H -V1b -V2b /2 width=3 by csx_appl_theta/
-#V1b #V2b #V1 #V2 #HV12 #HV12b #IHV12b #W1 #W2 #HW12 #V #T #H
-lapply (csx_appl_theta … HW12 … H) -H -HW12 #H
-lapply (csx_fwd_pair_sn … H) #HW1
-lapply (csx_fwd_flat_dx … H) #H1
-@csx_appl_simple_theq /2 width=3 by simple_flat/ -IHV12b -HW1 -H1 #X #H1 #H2
-elim (cpxs_fwd_theta_vector … (V2@V2b) … H1) -H1 /2 width=1 by liftv_cons/ -HV12b -HV12
-[ -H #H elim H2 -H2 //
-| -H2 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
-]
-qed.
-
-(* Basic_1: was just: sn3_appls_cast *)
-lemma csx_applv_cast: ∀h,o,G,L,Vs,W,T. ⦃G, L⦄ ⊢ ⬈*[h, o] ⒶVs.W → ⦃G, L⦄ ⊢ ⬈*[h, o] ⒶVs.T →
-                      ⦃G, L⦄ ⊢ ⬈*[h, o] ⒶVs.ⓝW.T.
-#h #o #G #L #Vs elim Vs -Vs /2 width=1 by csx_cast/
-#V #Vs #IHV #W #T #H1W #H1T
-lapply (csx_fwd_pair_sn … H1W) #HV
-lapply (csx_fwd_flat_dx … H1W) #H2W
-lapply (csx_fwd_flat_dx … H1T) #H2T
-@csx_appl_simple_theq /2 width=1 by applv_simple, simple_flat/ -IHV -HV -H2W -H2T
-#X #H #H0
-elim (cpxs_fwd_cast_vector … H) -H #H
-[ -H1W -H1T elim H0 -H0 //
-| -H1W -H0 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
-| -H1T -H0 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
-]
-qed.
-
-theorem csx_gcr: ∀h,o. gcr (cpx h o) (eq …) (csx h o) (csx h o).
-#h #o @mk_gcr //
-[ /3 width=1 by csx_applv_cnx/
-|2,3,6: /2 width=1 by csx_applv_beta, csx_applv_sort, csx_applv_cast/
-| /2 width=7 by csx_applv_delta/
-| #G #L #V1b #V2b #HV12b #a #V #T #H #HV
-  @(csx_applv_theta … HV12b) -HV12b
-  @csx_abbr //
-]
-qed.
index 6c76c4bc85e45e0e3ffa0325876face3f55e99ec..2ed89fa2f271bd6bcdd9a8f67c0d92a06b21b265 100644 (file)
@@ -1,4 +1,4 @@
 cpxs.ma cpxs_tdeq.ma cpxs_theq.ma cpxs_theq_vector.ma cpxs_drops.ma cpxs_lsubr.ma cpxs_lfpx.ma cpxs_cnx.ma cpxs_cpxs.ma
 lfpxs.ma lfpxs_fqup.ma lfpxs_cpxs.ma
-csx.ma csx_simple.ma csx_simple_theq.ma csx_drops.ma csx_lsubr.ma csx_gcp.ma csx_lfpx.ma csx_cnx.ma csx_cpxs.ma csx_csx.ma
-csx_vector.ma csx_cnx_vector.ma
+csx.ma csx_simple.ma csx_simple_theq.ma csx_drops.ma csx_lsubr.ma csx_gcp.ma csx_gcr.ma csx_lfpx.ma csx_cnx.ma csx_cpxs.ma csx_csx.ma
+csx_vector.ma csx_cnx_vector.ma csx_csx_vector.ma 
index d2d7a0e35bc5a01cc38de6a6f0a5739781c02fc9..c1b25de72761f70c15e48a2b0405f73fda24c115 100644 (file)
@@ -113,8 +113,8 @@ table {
              [ "lpxs ( ⦃?,?⦄ ⊢ ➡*[?,?] ? )" "lpxs_drop" + "lpxs_lleq" + "lpxs_aaa" + "lpxs_cpxs" + "lpxs_lpxs" * ]
              [ "cpxs_lreq" + "cpxs_lleq" + "cpxs_aaa" * ]
 *)
-             [ "csx_vector ( ⦃?,?⦄ ⊢ ⬈*[?,?] 𝐒⦃?⦄ )"  "csx_cnx_vector" * ]
-             [ "csx ( ⦃?,?⦄ ⊢ ⬈*[?,?] 𝐒⦃?⦄ )" "csx_simple" + "csx_simple_theq" + "csx_drops" + "csx_lsubr" + "csx_gcp" + "csx_lfpx" + "csx_cnx" + "csx_cpxs" + "csx_csx" * ]
+             [ "csx_vector ( ⦃?,?⦄ ⊢ ⬈*[?,?] 𝐒⦃?⦄ )" "csx_cnx_vector" + "csx_csx_vector" * ]
+             [ "csx ( ⦃?,?⦄ ⊢ ⬈*[?,?] 𝐒⦃?⦄ )" "csx_simple" + "csx_simple_theq" + "csx_drops" + "csx_lsubr" + "csx_gcp" + "csx_gcr" + "csx_lfpx" + "csx_cnx" + "csx_cpxs" + "csx_csx" * ]
              [ "lfpxs ( ⦃?,?⦄ ⊢ ⬈*[?,?] ? )" "lfpxs_length" + "lfpxs_fqup" + "lfpxs_cpxs" * ]
              [ "cpxs ( ⦃?,?⦄ ⊢ ? ⬈*[?] ? )" "cpxs_tdeq" + "cpxs_theq" + "cpxs_theq_vector" + "cpxs_drops" + "cpxs_lsubr" + "cpxs_lfpx" + "cpxs_cnx" + "cpxs_cpxs" * ] 
           }