]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda-delta/Basic-2/substitution/tps.ma
- the substitution lemma is proved!
[helm.git] / matita / matita / contribs / lambda-delta / Basic-2 / substitution / tps.ma
index 56d8f32118d77b893745692aa78b566470ab39c6..0acbc3d4fa195acfc1cdd1b68c3a48d3121c7664 100644 (file)
@@ -1,46 +1,44 @@
-(*
-    ||M||  This file is part of HELM, an Hypertextual, Electronic
-    ||A||  Library of Mathematics, developed at the Computer Science
-    ||T||  Department of the University of Bologna, Italy.
-    ||I||
-    ||T||
-    ||A||  This file is distributed under the terms of the
-    \   /  GNU General Public License Version 2
-     \ /
-      V_______________________________________________________________ *)
-
-include "lambda-delta/substitution/drop.ma".
-
-(* PARTIAL SUBSTITUTION ON TERMS ********************************************)
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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/drop.ma".
+
+(* PARALLEL SUBSTITUTION ON TERMS *******************************************)
 
 inductive tps: lenv → term → nat → nat → term → Prop ≝
-| tps_sort : ∀L,k,d,e. tps L (⋆k) d e (⋆k)
-| tps_lref : ∀L,i,d,e. tps L (#i) d e (#i)
-| tps_subst: ∀L,K,V,U1,U2,i,d,e.
-             d ≤ i → i < d + e →
-             ↓[0, i] L ≡ K. 𝕓{Abbr} V → tps K V 0 (d + e - i - 1) U1 →
-             ↑[0, i + 1] U1 ≡ U2 → tps L (#i) d e U2
+| tps_atom : ∀L,I,d,e. tps L (𝕒{I}) d e (𝕒{I})
+| tps_subst: ∀L,K,V,W,i,d,e. d ≤ i → i < d + e →
+             ↓[0, i] L ≡ K. 𝕓{Abbr} V → ↑[0, i + 1] V ≡ W → tps L (#i) d e W
 | tps_bind : ∀L,I,V1,V2,T1,T2,d,e.
-             tps L V1 d e V2 → tps (L. 𝕓{I} V1) T1 (d + 1) e T2 →
+             tps L V1 d e V2 → tps (L. 𝕓{I} V2) T1 (d + 1) e T2 →
              tps L (𝕓{I} V1. T1) d e (𝕓{I} V2. T2)
 | tps_flat : ∀L,I,V1,V2,T1,T2,d,e.
              tps L V1 d e V2 → tps L T1 d e T2 →
              tps L (𝕗{I} V1. T1) d e (𝕗{I} V2. T2)
 .
 
-interpretation "partial telescopic substritution"
+interpretation "parallel substritution (term)"
    'PSubst L T1 d e T2 = (tps L T1 d e T2).
 
 (* Basic properties *********************************************************)
 
-lemma tps_leq_repl: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ≫ T2 →
-                    ∀L2. L1 [d, e] ≈ L2 → L2 ⊢ T1 [d, e] ≫ T2.
+lemma tps_leq_repl_dx: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ≫ T2 →
+                       ∀L2. L1 [d, e] ≈ L2 → L2 ⊢ T1 [d, e] ≫ T2.
 #L1 #T1 #T2 #d #e #H elim H -H L1 T1 T2 d e
 [ //
-| //
-| #L1 #K1 #V #V1 #V2 #i #d #e #Hdi #Hide #HLK1 #_ #HV12 #IHV12 #L2 #HL12
-  elim (drop_leq_drop1 … HL12 … HLK1 ? ?) -HL12 HLK1 // #K2 #HK12 #HLK2
-  @tps_subst [4,5,6,8: // |1,2,3: skip | /2/ ] (**) (* /3 width=6/ is too slow *)
+| #L1 #K1 #V #W #i #d #e #Hdi #Hide #HLK1 #HVW #L2 #HL12
+  elim (drop_leq_drop1 … HL12 … HLK1 ? ?) -HL12 HLK1 // /2/
 | /4/
 | /3/
 ]
@@ -56,11 +54,9 @@ lemma tps_weak: ∀L,T1,T2,d1,e1. L ⊢ T1 [d1, e1] ≫ T2 →
                 L ⊢ T1 [d2, e2] ≫ T2.
 #L #T1 #T #d1 #e1 #H elim H -L T1 T d1 e1
 [ //
-| //
-| #L #K #V #V1 #V2 #i #d1 #e1 #Hid1 #Hide1 #HLK #_ #HV12 #IHV12 #d2 #e2 #Hd12 #Hde12
+| #L #K #V #W #i #d1 #e1 #Hid1 #Hide1 #HLK #HVW #d2 #e2 #Hd12 #Hde12
   lapply (transitive_le … Hd12 … Hid1) -Hd12 Hid1 #Hid2
-  lapply (lt_to_le_to_lt … Hide1 … Hde12) -Hide1 #Hide2
-  @tps_subst [4,5,6,8: // |1,2,3: skip | @IHV12 /2/ ] (**) (* /4 width=6/ is too slow *)
+  lapply (lt_to_le_to_lt … Hide1 … Hde12) -Hide1 /2/
 | /4/
 | /4/
 ]
@@ -70,13 +66,10 @@ lemma tps_weak_top: ∀L,T1,T2,d,e.
                     L ⊢ T1 [d, e] ≫ T2 → L ⊢ T1 [d, |L| - d] ≫ T2.
 #L #T1 #T #d #e #H elim H -L T1 T d e
 [ //
-| //
-| #L #K #V #V1 #V2 #i #d #e #Hdi #_ #HLK #_ #HV12 #IHV12
+| #L #K #V #W #i #d #e #Hdi #_ #HLK #HVW
   lapply (drop_fwd_drop2_length … HLK) #Hi
   lapply (le_to_lt_to_lt … Hdi Hi) #Hd
-  lapply (plus_minus_m_m_comm (|L|) d ?) [ /2/ ] -Hd #Hd
-  lapply (drop_fwd_O1_length … HLK) normalize #HKL
-  lapply (tps_weak … IHV12 0 (|L| - i - 1) ? ?) -IHV12 // -HKL /2 width=6/
+  lapply (plus_minus_m_m_comm (|L|) d ?) /2/
 | normalize /2/
 | /2/
 ]
@@ -87,42 +80,82 @@ lemma tps_weak_all: ∀L,T1,T2,d,e.
 #L #T1 #T #d #e #HT12
 lapply (tps_weak … HT12 0 (d + e) ? ?) -HT12 // #HT12
 lapply (tps_weak_top … HT12) //
-qed.     
+qed.
+
+lemma tps_split_up: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → ∀i. d ≤ i → i ≤ d + e →
+                    ∃∃T. L ⊢ T1 [d, i - d] ≫ T & L ⊢ T [i, d + e - i] ≫ T2.
+#L #T1 #T2 #d #e #H elim H -L T1 T2 d e
+[ /2/
+| #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hdj #Hjde
+  elim (lt_or_ge i j)
+  [ -Hide Hjde;
+    >(plus_minus_m_m_comm j d) in ⊢ (% → ?) // -Hdj /3/
+  | -Hdi Hdj; #Hid
+    generalize in match Hide -Hide (**) (* rewriting in the premises, rewrites in the goal too *)
+    >(plus_minus_m_m_comm … Hjde) in ⊢ (% → ?) -Hjde /4/
+  ]
+| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide
+  elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2
+  elim (IHT12 (i + 1) ? ?) -IHT12 [2: /2 by arith4/ |3: /2/ ] (* just /2/ is too slow *)
+  -Hdi Hide >arith_c1 >arith_c1x #T #HT1 #HT2
+  lapply (tps_leq_repl_dx … HT1 (L. 𝕓{I} V) ?) -HT1 /3 width=5/
+| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide
+  elim (IHV12 i ? ?) -IHV12 // elim (IHT12 i ? ?) -IHT12 //
+  -Hdi Hide /3 width=5/
+]
+qed.
 
 (* Basic inversion lemmas ***************************************************)
 
-lemma tps_inv_lref1_aux: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → ∀i. T1 = #i →
-                         T2 = #i ∨ 
-                         ∃∃K,V1,V2,i. d ≤ i & i < d + e &
-                                      ↓[O, i] L ≡ K. 𝕓{Abbr} V1 &
-                                      K ⊢ V1 [O, d + e - i - 1] ≫ V2 &
-                                      ↑[O, i + 1] V2 ≡ T2.
+fact tps_inv_atom1_aux: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → ∀I. T1 = 𝕒{I} →
+                        T2 = 𝕒{I} ∨
+                        ∃∃K,V,i. d ≤ i & i < d + e &
+                                 ↓[O, i] L ≡ K. 𝕓{Abbr} V &
+                                 ↑[O, i + 1] V ≡ T2 &
+                                 I = LRef i.
 #L #T1 #T2 #d #e * -L T1 T2 d e
-[ #L #k #d #e #i #H destruct
-| /2/
-| #L #K #V1 #V2 #T2 #i #d #e #Hdi #Hide #HLK #HV12 #HVT2 #j #H destruct -i /3 width=9/
-| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct
-| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct
+[ #L #I #d #e #J #H destruct -I /2/
+| #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #I #H destruct -I /3 width=8/
+| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct
+| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct
 ]
 qed.
 
-lemma tps_inv_lref1: ∀L,T2,i,d,e. L ⊢ #i [d, e] ≫ T2 →
-                     T2 = #i ∨ 
-                     ∃∃K,V1,V2,i. d ≤ i & i < d + e &
-                                  ↓[O, i] L ≡ K. 𝕓{Abbr} V1 &
-                                  K ⊢ V1 [O, d + e - i - 1] ≫ V2 &
-                                  ↑[O, i + 1] V2 ≡ T2.
+lemma tps_inv_atom1: ∀L,T2,I,d,e. L ⊢ 𝕒{I} [d, e] ≫ T2 →
+                     T2 = 𝕒{I} ∨
+                     ∃∃K,V,i. d ≤ i & i < d + e &
+                              ↓[O, i] L ≡ K. 𝕓{Abbr} V &
+                              ↑[O, i + 1] V ≡ T2 &
+                              I = LRef i.
 /2/ qed.
 
-lemma tps_inv_bind1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ≫ U2 →
-                         ∀I,V1,T1. U1 = 𝕓{I} V1. T1 →
-                         ∃∃V2,T2. L ⊢ V1 [d, e] ≫ V2 & 
-                                  L. 𝕓{I} V1 ⊢ T1 [d + 1, e] ≫ T2 &
-                                  U2 =  𝕓{I} V2. T2.
+
+(* Basic-1: was: subst1_gen_sort *)
+lemma tps_inv_sort1: ∀L,T2,k,d,e. L ⊢ ⋆k [d, e] ≫ T2 → T2 = ⋆k.
+#L #T2 #k #d #e #H
+elim (tps_inv_atom1 … H) -H //
+* #K #V #i #_ #_ #_ #_ #H destruct
+qed.
+
+(* Basic-1: was: subst1_gen_lref *)
+lemma tps_inv_lref1: ∀L,T2,i,d,e. L ⊢ #i [d, e] ≫ T2 →
+                     T2 = #i ∨
+                     ∃∃K,V. d ≤ i & i < d + e &
+                            ↓[O, i] L ≡ K. 𝕓{Abbr} V &
+                            ↑[O, i + 1] V ≡ T2.
+#L #T2 #i #d #e #H
+elim (tps_inv_atom1 … H) -H /2/
+* #K #V #j #Hdj #Hjde #HLK #HVT2 #H destruct -i /3/
+qed.
+
+fact tps_inv_bind1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ≫ U2 →
+                        ∀I,V1,T1. U1 = 𝕓{I} V1. T1 →
+                        ∃∃V2,T2. L ⊢ V1 [d, e] ≫ V2 & 
+                                 L. 𝕓{I} V2 ⊢ T1 [d + 1, e] ≫ T2 &
+                                 U2 =  𝕓{I} V2. T2.
 #d #e #L #U1 #U2 * -d e L U1 U2
 [ #L #k #d #e #I #V1 #T1 #H destruct
-| #L #i #d #e #I #V1 #T1 #H destruct
-| #L #K #V #U1 #U2 #i #d #e #_ #_ #_ #_ #_ #I #V1 #T1 #H destruct
+| #L #K #V #W #i #d #e #_ #_ #_ #_ #I #V1 #T1 #H destruct
 | #L #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #I #V #T #H destruct /2 width=5/
 | #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct
 ]
@@ -130,18 +163,17 @@ qed.
 
 lemma tps_inv_bind1: ∀d,e,L,I,V1,T1,U2. L ⊢ 𝕓{I} V1. T1 [d, e] ≫ U2 →
                      ∃∃V2,T2. L ⊢ V1 [d, e] ≫ V2 & 
-                              L. 𝕓{I} V1 ⊢ T1 [d + 1, e] ≫ T2 &
+                              L. 𝕓{I} V2 ⊢ T1 [d + 1, e] ≫ T2 &
                               U2 =  𝕓{I} V2. T2.
 /2/ qed.
 
-lemma tps_inv_flat1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ≫ U2 →
-                         ∀I,V1,T1. U1 = 𝕗{I} V1. T1 →
-                         ∃∃V2,T2. L ⊢ V1 [d, e] ≫ V2 & L ⊢ T1 [d, e] ≫ T2 &
-                                  U2 =  𝕗{I} V2. T2.
+fact tps_inv_flat1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ≫ U2 →
+                        ∀I,V1,T1. U1 = 𝕗{I} V1. T1 →
+                        ∃∃V2,T2. L ⊢ V1 [d, e] ≫ V2 & L ⊢ T1 [d, e] ≫ T2 &
+                                 U2 =  𝕗{I} V2. T2.
 #d #e #L #U1 #U2 * -d e L U1 U2
 [ #L #k #d #e #I #V1 #T1 #H destruct
-| #L #i #d #e #I #V1 #T1 #H destruct
-| #L #K #V #U1 #U2 #i #d #e #_ #_ #_ #_ #_ #I #V1 #T1 #H destruct
+| #L #K #V #W #i #d #e #_ #_ #_ #_ #I #V1 #T1 #H destruct
 | #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct
 | #L #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #I #V #T #H destruct /2 width=5/
 ]
@@ -151,3 +183,27 @@ lemma tps_inv_flat1: ∀d,e,L,I,V1,T1,U2. L ⊢ 𝕗{I} V1. T1 [d, e] ≫ U2 →
                      ∃∃V2,T2. L ⊢ V1 [d, e] ≫ V2 & L ⊢ T1 [d, e] ≫ T2 &
                               U2 =  𝕗{I} V2. T2.
 /2/ qed.
+
+fact tps_inv_refl_O2_aux: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → e = 0 → T1 = T2.
+#L #T1 #T2 #d #e #H elim H -H L T1 T2 d e
+[ //
+| #L #K #V #W #i #d #e #Hdi #Hide #_ #_ #H destruct -e;
+  lapply (le_to_lt_to_lt … Hdi … Hide) -Hdi Hide <plus_n_O #Hdd
+  elim (lt_refl_false … Hdd)
+| /3/
+| /3/
+]
+qed.
+
+lemma tps_inv_refl_O2: ∀L,T1,T2,d. L ⊢ T1 [d, 0] ≫ T2 → T1 = T2.
+/2 width=6/ qed.
+
+(* Basic-1: removed theorems 23:
+            subst0_gen_sort subst0_gen_lref subst0_gen_head subst0_gen_lift_lt
+            subst0_gen_lift_false subst0_gen_lift_ge subst0_refl subst0_trans
+                 subst0_lift_lt subst0_lift_ge subst0_lift_ge_S subst0_lift_ge_s
+            subst0_subst0 subst0_subst0_back subst0_weight_le subst0_weight_lt
+                 subst0_confluence_neq subst0_confluence_eq subst0_tlt_head
+            subst0_confluence_lift subst0_tlt
+            subst1_head subst1_gen_head  
+*)