From: Ferruccio Guidi Date: Mon, 8 Aug 2011 22:41:37 +0000 (+0000) Subject: - tps_tpr closed! (substitution is a reduction) X-Git-Tag: make_still_working~2341 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=9271b3ca211007ca5ffac1e7644ebc02b0689d6e;p=helm.git - tps_tpr closed! (substitution is a reduction) - some refactoring --- diff --git a/matita/matita/lib/lambda-delta/Makefile b/matita/matita/lib/lambda-delta/Makefile index 648904368..8feb040db 100644 --- a/matita/matita/lib/lambda-delta/Makefile +++ b/matita/matita/lib/lambda-delta/Makefile @@ -3,7 +3,7 @@ XOA_DIR = ../../../components/binaries/xoa XOA = xoa.native CONF = xoa.conf.xml -TARGETS = xoa_natation.ma xoa_defs.ma +TARGETS = xoa_natation.ma xoa.ma all: $(TARGETS) diff --git a/matita/matita/lib/lambda-delta/ground.ma b/matita/matita/lib/lambda-delta/ground.ma index 58d86aa5a..92a3f8167 100644 --- a/matita/matita/lib/lambda-delta/ground.ma +++ b/matita/matita/lib/lambda-delta/ground.ma @@ -74,12 +74,13 @@ lemma le_plus_minus_comm: ∀n,m,p. p ≤ m → m + n - p = m - p + n. >(commutative_plus p) (le_O_to_eq_O … H) -H // -| #m #IHm #p elim p -p // - #p #_ #n #Hpm (le_O_to_eq_O … H) -H // +| #b #IHb #c elim c -c // + #c #_ #a #Hcb + lapply (le_S_S_to_le … Hcb) -Hcb #Hcb + (plus_minus_m_m e 1) /2/ qed. -lemma drop_fwd_drop2: ∀L1,I2,K2,V2,e. ↓[O, e] L1 ≡ K2. 𝕓{I2} V2 → - ↓[O, e + 1] L1 ≡ K2. -#L1 elim L1 -L1 -[ #I2 #K2 #V2 #e #H lapply (drop_inv_sort1 … H) -H #H destruct -| #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #e #H - elim (drop_inv_O1 … H) -H * #He #H - [ -IHL1; destruct -e K2 I2 V2 /2/ - | @drop_drop >(plus_minus_m_m e 1) /2/ - ] -] -qed. - lemma drop_leq_drop1: ∀L1,L2,d,e. L1 [d, e] ≈ L2 → ∀I,K1,V,i. ↓[0, i] L1 ≡ K1. 𝕓{I} V → d ≤ i → i < d + e → @@ -145,3 +133,41 @@ lemma drop_leq_drop1: ∀L1,L2,d,e. L1 [d, e] ≈ L2 → elim (IHL12 … HLK1 ? ?) -IHL12 HLK1 [2: /2/ |3: /2/ ] -Hdi Hide >arith_g1 // /3/ ] qed. + +(* Basic forvard lemmas *****************************************************) + +lemma drop_fwd_drop2: ∀L1,I2,K2,V2,e. ↓[O, e] L1 ≡ K2. 𝕓{I2} V2 → + ↓[O, e + 1] L1 ≡ K2. +#L1 elim L1 -L1 +[ #I2 #K2 #V2 #e #H lapply (drop_inv_sort1 … H) -H #H destruct +| #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #e #H + elim (drop_inv_O1 … H) -H * #He #H + [ -IHL1; destruct -e K2 I2 V2 /2/ + | @drop_drop >(plus_minus_m_m e 1) /2/ + ] +] +qed. + +lemma drop_fwd_drop2_length: ∀L1,I2,K2,V2,e. + ↓[0, e] L1 ≡ K2. 𝕓{I2} V2 → e < |L1|. +#L1 elim L1 -L1 +[ #I2 #K2 #V2 #e #H lapply (drop_inv_sort1 … H) -H #H destruct +| #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #e #H + elim (drop_inv_O1 … H) -H * #He #H + [ -IHL1; destruct -e K2 I2 V2 // + | lapply (IHL1 … H) -IHL1 H #HeK1 whd in ⊢ (? ? %) /2/ + ] +] +qed. + +lemma drop_fwd_O1_length: ∀L1,L2,e. ↓[0, e] L1 ≡ L2 → |L2| = |L1| - e. +#L1 elim L1 -L1 +[ #L2 #e #H >(drop_inv_sort1 … H) -H // +| #K1 #I1 #V1 #IHL1 #L2 #e #H + elim (drop_inv_O1 … H) -H * #He #H + [ -IHL1; destruct -e L2 // + | lapply (IHL1 … H) -IHL1 H #H >H -H; normalize + >minus_le_minus_minus_comm // + ] +] +qed. diff --git a/matita/matita/lib/lambda-delta/substitution/drop_drop.ma b/matita/matita/lib/lambda-delta/substitution/drop_drop.ma index 96fd0db51..135db64ca 100644 --- a/matita/matita/lib/lambda-delta/substitution/drop_drop.ma +++ b/matita/matita/lib/lambda-delta/substitution/drop_drop.ma @@ -76,7 +76,7 @@ lemma drop_trans_le: ∀d1,e1,L1,L. ↓[d1, e1] L1 ≡ L → [ -He2d IHL12 #H1 #H2 destruct -e2 L /3 width=5/ | -HL12 HV12 #He2 #HL2 elim (IHL12 … HL2 ?) -IHL12 HL2 L2 - [ minus_le_minus_minus_comm // /3/ | /2/ ] ] ] qed. diff --git a/matita/matita/lib/lambda-delta/substitution/leq.ma b/matita/matita/lib/lambda-delta/substitution/leq.ma index 3c7981566..b0e28d48e 100644 --- a/matita/matita/lib/lambda-delta/substitution/leq.ma +++ b/matita/matita/lib/lambda-delta/substitution/leq.ma @@ -9,7 +9,7 @@ \ / V_______________________________________________________________ *) -include "lambda-delta/syntax/lenv.ma". +include "lambda-delta/syntax/length.ma". (* LOCAL ENVIRONMENT EQUALITY ***********************************************) @@ -37,12 +37,16 @@ lemma leq_sym: ∀L1,L2,d,e. L1 [d, e] ≈ L2 → L2 [d, e] ≈ L1. #L1 #L2 #d #e #H elim H -H L1 L2 d e /2/ qed. -lemma leq_skip_lt: ∀L1,L2,d,e. leq L1 (d - 1) e L2 → 0 < d → +lemma leq_skip_lt: ∀L1,L2,d,e. L1 [d - 1, e] ≈ L2 → 0 < d → ∀I1,I2,V1,V2. L1. 𝕓{I1} V1 [d, e] ≈ L2. 𝕓{I2} V2. #L1 #L2 #d #e #HL12 #Hd >(plus_minus_m_m d 1) /2/ qed. +lemma leq_fwd_length: ∀L1,L2,d,e. L1 [d, e] ≈ L2 → |L1| = |L2|. +#L1 #L2 #d #e #H elim H -H L1 L2 d e; normalize // +qed. + (* Basic inversion lemmas ***************************************************) lemma leq_inv_sort1_aux: ∀L1,L2,d,e. L1 [d, e] ≈ L2 → L1 = ⋆ → L2 = ⋆. diff --git a/matita/matita/lib/lambda-delta/substitution/pts.ma b/matita/matita/lib/lambda-delta/substitution/pts.ma deleted file mode 100644 index ac962355e..000000000 --- a/matita/matita/lib/lambda-delta/substitution/pts.ma +++ /dev/null @@ -1,107 +0,0 @@ -(* - ||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 TELESCOPIC SUBSTITUTION ******************************************) - -inductive pts: lenv → term → nat → nat → term → Prop ≝ -| pts_sort : ∀L,k,d,e. pts L (⋆k) d e (⋆k) -| pts_lref : ∀L,i,d,e. pts L (#i) d e (#i) -| pts_subst: ∀L,K,V,U1,U2,i,d,e. - d ≤ i → i < d + e → - ↓[0, i] L ≡ K. 𝕓{Abbr} V → pts K V 0 (d + e - i - 1) U1 → - ↑[0, i + 1] U1 ≡ U2 → pts L (#i) d e U2 -| pts_bind : ∀L,I,V1,V2,T1,T2,d,e. - pts L V1 d e V2 → pts (L. 𝕓{I} V1) T1 (d + 1) e T2 → - pts L (𝕓{I} V1. T1) d e (𝕓{I} V2. T2) -| pts_flat : ∀L,I,V1,V2,T1,T2,d,e. - pts L V1 d e V2 → pts L T1 d e T2 → - pts L (𝕗{I} V1. T1) d e (𝕗{I} V2. T2) -. - -interpretation "partial telescopic substritution" - 'PSubst L T1 d e T2 = (pts L T1 d e T2). - -(* Basic properties *********************************************************) - -lemma pts_leq_repl: ∀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 - @pts_subst [4,5,6,8: // |1,2,3: skip | /2/ ] (**) (* /3 width=6/ is too slow *) -| /4/ -| /3/ -] -qed. - -lemma pts_refl: ∀T,L,d,e. L ⊢ T [d, e] ≫ T. -#T elim T -T // -#I elim I -I /2/ -qed. - -lemma pts_weak: ∀L,T1,T2,d1,e1. L ⊢ T1 [d1, e1] ≫ T2 → - ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → - 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 - lapply (transitive_le … Hd12 … Hid1) -Hd12 Hid1 #Hid2 - lapply (lt_to_le_to_lt … Hide1 … Hde12) -Hide1 #Hide2 - @pts_subst [4,5,6,8: // |1,2,3: skip | @IHV12 /2/ ] (**) (* /4 width=6/ is too slow *) -| /4/ -| /4/ -] -qed. - -(* Basic inversion lemmas ***************************************************) - -lemma pts_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. -#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 #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 -] -qed. - -lemma pts_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 & - U2 = 𝕓{I} V2. T2. -/2/ qed. - -lemma pts_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 #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/ -] -qed. - -lemma pts_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. diff --git a/matita/matita/lib/lambda-delta/substitution/pts_lift.ma b/matita/matita/lib/lambda-delta/substitution/pts_lift.ma deleted file mode 100644 index 8cd477738..000000000 --- a/matita/matita/lib/lambda-delta/substitution/pts_lift.ma +++ /dev/null @@ -1,174 +0,0 @@ -(* - ||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/lift_lift.ma". -include "lambda-delta/substitution/drop_drop.ma". -include "lambda-delta/substitution/pts.ma". - -(* PARTIAL TELESCOPIC SUBSTITUTION ******************************************) - -(* Relocation properties ****************************************************) - -lemma pts_lift_le: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ≫ T2 → - ∀L,U1,U2,d,e. ↓[d, e] L ≡ K → - ↑[d, e] T1 ≡ U1 → ↑[d, e] T2 ≡ U2 → - dt + et ≤ d → - L ⊢ U1 [dt, et] ≫ U2. -#K #T1 #T2 #dt #et #H elim H -H K T1 T2 dt et -[ #K #k #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ - lapply (lift_mono … H1 … H2) -H1 H2 #H destruct -U1 // -| #K #i #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ - lapply (lift_mono … H1 … H2) -H1 H2 #H destruct -U1 // -| #K #KV #V #V1 #V2 #i #dt #et #Hdti #Hidet #HKV #_ #HV12 #IHV12 #L #U1 #U2 #d #e #HLK #H #HVU2 #Hdetd - lapply (lt_to_le_to_lt … Hidet … Hdetd) #Hid - lapply (lift_inv_lref1_lt … H … Hid) -H #H destruct -U1; - elim (lift_trans_ge … HV12 … HVU2 ?) -HV12 HVU2 V2 // plus_plus_comm_23 #HV1U2 - lapply (drop_trans_ge_comm … HLK … HKV ?) -HLK HKV K // -Hid #HLKV - @pts_subst [4,5: /2/ |6,7,8: // |1,2,3: skip ] (**) (* /3 width=8/ is too slow *) -| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt - elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct -U1 U2; - @pts_bind [ /2 width=5/ | /3 width=5/ ] (**) (* explicit constructor *) -| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt - elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct -U1 U2; - /3 width=5/ -] -qed. - -lemma pts_inv_lift1_le: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 → - ∀K,d,e. ↓[d, e] L ≡ K → ∀T1. ↑[d, e] T1 ≡ U1 → - dt + et ≤ d → - ∃∃T2. K ⊢ T1 [dt, et] ≫ T2 & ↑[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H elim H -H L U1 U2 dt et -[ #L #k #dt #et #K #d #e #_ #T1 #H #_ - lapply (lift_inv_sort2 … H) -H #H destruct -T1 /2/ -| #L #i #dt #et #K #d #e #_ #T1 #H #_ - elim (lift_inv_lref2 … H) -H * #Hid #H destruct -T1 /3/ -| #L #KV #V #V1 #V2 #i #dt #et #Hdti #Hidet #HLKV #_ #HV12 #IHV12 #K #d #e #HLK #T1 #H #Hdetd - lapply (lt_to_le_to_lt … Hidet … Hdetd) #Hid - lapply (lift_inv_lref2_lt … H … Hid) -H #H destruct -T1; - elim (drop_conf_lt … HLK … HLKV ?) -HLK HLKV L // #L #W #HKL #HKVL #HWV - elim (IHV12 … HKVL … HWV ?) -HKVL HWV /2/ -Hdetd #W1 #HW1 #HWV1 - elim (lift_trans_le … HWV1 … HV12 ?) -HWV1 HV12 V1 // >arith_a2 /3 width=6/ -| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd - elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct -X; - elim (IHV12 … HLK … HWV1 ?) -IHV12 // - elim (IHU12 … HTU1 ?) -IHU12 HTU1 [3: /2/ |4: @drop_skip // |2: skip ] -HLK HWV1 Hdetd /3 width=5/ (**) (* just /3 width=5/ is too slow *) -| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd - elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct -X; - elim (IHV12 … HLK … HWV1 ?) -IHV12 HWV1 // - elim (IHU12 … HLK … HTU1 ?) -IHU12 HLK HTU1 // /3 width=5/ -] -qed. - -lemma pts_inv_lift1_ge: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 → - ∀K,d,e. ↓[d, e] L ≡ K → ∀T1. ↑[d, e] T1 ≡ U1 → - d + e ≤ dt → - ∃∃T2. K ⊢ T1 [dt - e, et] ≫ T2 & ↑[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H elim H -H L U1 U2 dt et -[ #L #k #dt #et #K #d #e #_ #T1 #H #_ - lapply (lift_inv_sort2 … H) -H #H destruct -T1 /2/ -| #L #i #dt #et #K #d #e #_ #T1 #H #_ - elim (lift_inv_lref2 … H) -H * #Hid #H destruct -T1 /3/ -| #L #KV #V #V1 #V2 #i #dt #et #Hdti #Hidet #HLKV #HV1 #HV12 #_ #K #d #e #HLK #T1 #H #Hdedt - lapply (transitive_le … Hdedt … Hdti) #Hdei - lapply (plus_le_weak … Hdedt) -Hdedt #Hedt - lapply (plus_le_weak … Hdei) #Hei - <(arith_h1 ? ? ? e ? ?) in HV1 // #HV1 - lapply (lift_inv_lref2_ge … H … Hdei) -H #H destruct -T1; - lapply (drop_conf_ge … HLK … HLKV ?) -HLK HLKV L // #HKV - elim (lift_split … HV12 d (i - e + 1) ? ? ?) -HV12; [2,3,4: normalize /2/ ] -Hdei >arith_e2 // #V0 #HV10 #HV02 - @ex2_1_intro - [2: @pts_subst [4: /2/ |6,7,8: // |1,2,3: skip |5: @arith5 // ] - |1: skip - | // - ] (**) (* explicitc constructors *) -| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd - elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct -X; - lapply (plus_le_weak … Hdetd) #Hedt - elim (IHV12 … HLK … HWV1 ?) -IHV12 // #W2 #HW12 #HWV2 - elim (IHU12 … HTU1 ?) -IHU12 HTU1 [4: @drop_skip // |2: skip |3: /2/ ] - IHV12 // >IHT12 // -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX - elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #H destruct -X - >IHV12 // >IHT12 // -] -qed. -(* - Theorem subst0_gen_lift_ge : (u,t1,x:?; i,h,d:?) (subst0 i u (lift h d t1) x) -> - (le (plus d h) i) -> - (EX t2 | x = (lift h d t2) & (subst0 (minus i h) u t1 t2)). - - Theorem subst0_gen_lift_rev_ge: (t1,v,u2:?; i,h,d:?) - (subst0 i v t1 (lift h d u2)) -> - (le (plus d h) i) -> - (EX u1 | (subst0 (minus i h) v u1 u2) & - t1 = (lift h d u1) - ). - - - Theorem subst0_gen_lift_rev_lelt: (t1,v,u2:?; i,h,d:?) - (subst0 i v t1 (lift h d u2)) -> - (le d i) -> (lt i (plus d h)) -> - (EX u1 | t1 = (lift (minus (plus d h) (S i)) (S i) u1)). -*) diff --git a/matita/matita/lib/lambda-delta/substitution/pts_pts.ma b/matita/matita/lib/lambda-delta/substitution/pts_pts.ma deleted file mode 100644 index 0461a1823..000000000 --- a/matita/matita/lib/lambda-delta/substitution/pts_pts.ma +++ /dev/null @@ -1,47 +0,0 @@ -(* - ||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/pts_split.ma". - -(* PARTIAL TELESCOPIC SUBSTITUTION ******************************************) - -(* Main properties **********************************************************) - -lemma pts_trans: ∀L,T1,T,d,e. L ⊢ T1 [d, e] ≫ T → ∀T2. L ⊢ T [d, e] ≫ T2 → - L ⊢ T1 [d, e] ≫ T2. -#L #T1 #T #d #e #H elim H -L T1 T d e -[ // -| // -| #L #K #V #V1 #V2 #i #d #e #Hdi #Hide #HLK #_ #HV12 #IHV12 #T2 #HVT2 - lapply (drop_fwd_drop2 … HLK) #H - elim (pts_inv_lift1_up … HVT2 … H … HV12 ? ? ?) -HVT2 H HV12 // normalize [2: /2/ ] #V #HV1 #HVT2 - @pts_subst [4,5,6,8: // |1,2,3: skip | /2/ ] -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX - elim (pts_inv_bind1 … HX) -HX #V #T #HV2 #HT2 #HX destruct -X; - @pts_bind /2/ @IHT12 @(pts_leq_repl … HT2) /2/ -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX - elim (pts_inv_flat1 … HX) -HX #V #T #HV2 #HT2 #HX destruct -X /3/ -] -qed. - -axiom pts_conf: ∀L,T0,d,e,T1. L ⊢ T0 [d, e] ≫ T1 → ∀T2. L ⊢ T0 [d, e] ≫ T2 → - ∃∃T. L ⊢ T1 [d, e] ≫ T & L ⊢ T2 [d, e] ≫ T. - -(* - Theorem subst0_subst0: (t1,t2,u2:?; j:?) (subst0 j u2 t1 t2) -> - (u1,u:?; i:?) (subst0 i u u1 u2) -> - (EX t | (subst0 j u1 t1 t) & (subst0 (S (plus i j)) u t t2)). - - Theorem subst0_subst0_back: (t1,t2,u2:?; j:?) (subst0 j u2 t1 t2) -> - (u1,u:?; i:?) (subst0 i u u2 u1) -> - (EX t | (subst0 j u1 t1 t) & (subst0 (S (plus i j)) u t2 t)). - -*) diff --git a/matita/matita/lib/lambda-delta/substitution/pts_split.ma b/matita/matita/lib/lambda-delta/substitution/pts_split.ma deleted file mode 100644 index 72da008ea..000000000 --- a/matita/matita/lib/lambda-delta/substitution/pts_split.ma +++ /dev/null @@ -1,58 +0,0 @@ -(* - ||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/pts_lift.ma". - -(* PARTIAL TELESCOPIC SUBSTITUTION ******************************************) - -(* Split properties *********************************************************) - -lemma pts_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/ -| /2/ -| #L #K #V #V1 #V2 #i #d #e #Hdi #Hide #HLK #HV1 #HV12 #IHV12 #j #Hdj #Hjde - elim (lt_or_ge i j) #Hij - [ -HV1 Hide; - lapply (drop_fwd_drop2 … HLK) #HLK' - elim (IHV12 (j - i - 1) ? ?) -IHV12; normalize /2/ -Hjde arith_b2 // #W1 #HVW1 #HWV1 - generalize in match HVW1 generalize in match Hij -HVW1 (**) (* rewriting in the premises, rewrites in the goal too *) - >(plus_minus_m_m_comm … Hdj) in ⊢ (% → % → ?) -Hdj #Hij' #HVW1 - elim (lift_total W1 0 (i + 1)) #W2 #HW12 - lapply (pts_lift_ge … HWV1 … HLK' HW12 HV12 ?) -HWV1 HLK' HV12 // >arith_a2 /3 width=6/ - | -IHV12 Hdi Hdj; - generalize in match HV1 generalize in match Hide -HV1 Hide (**) (* rewriting in the premises, rewrites in the goal too *) - >(plus_minus_m_m_comm … Hjde) in ⊢ (% → % → ?) -Hjde #Hide #HV1 - @ex2_1_intro [2: @pts_lref |1: skip | /2 width=6/ ] (**) (* /3 width=6 is too slow *) - ] -| #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 - @ex2_1_intro [2,3: @pts_bind | skip ] (**) (* explicit constructors *) - [3: @HV1 |4: @HT1 |5: // |1,2: skip | /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. - -lemma pts_inv_lift1_up: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 → - ∀K,d,e. ↓[d, e] L ≡ K → ∀T1. ↑[d, e] T1 ≡ U1 → - d ≤ dt → dt ≤ d + e → d + e ≤ dt + et → - ∃∃T2. K ⊢ T1 [d, dt + et - (d + e)] ≫ T2 & ↑[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #HU12 #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet -elim (pts_split_up … HU12 (d + e) ? ?) -HU12 // -Hdedet #U #HU1 #HU2 -lapply (pts_weak … HU1 d e ? ?) -HU1 // plus_plus_comm_23 #HV1U2 + lapply (drop_trans_ge_comm … HLK … HKV ?) -HLK HKV K // -Hid #HLKV + @tps_subst [4,5: /2/ |6,7,8: // |1,2,3: skip ] (**) (* /3 width=8/ is too slow *) +| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt + elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct -U1 U2; + @tps_bind [ /2 width=5/ | /3 width=5/ ] (**) (* explicit constructor *) +| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt + elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct -U1 U2; + /3 width=5/ +] +qed. + +lemma tps_inv_lift1_le: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 → + ∀K,d,e. ↓[d, e] L ≡ K → ∀T1. ↑[d, e] T1 ≡ U1 → + dt + et ≤ d → + ∃∃T2. K ⊢ T1 [dt, et] ≫ T2 & ↑[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #H elim H -H L U1 U2 dt et +[ #L #k #dt #et #K #d #e #_ #T1 #H #_ + lapply (lift_inv_sort2 … H) -H #H destruct -T1 /2/ +| #L #i #dt #et #K #d #e #_ #T1 #H #_ + elim (lift_inv_lref2 … H) -H * #Hid #H destruct -T1 /3/ +| #L #KV #V #V1 #V2 #i #dt #et #Hdti #Hidet #HLKV #_ #HV12 #IHV12 #K #d #e #HLK #T1 #H #Hdetd + lapply (lt_to_le_to_lt … Hidet … Hdetd) #Hid + lapply (lift_inv_lref2_lt … H … Hid) -H #H destruct -T1; + elim (drop_conf_lt … HLK … HLKV ?) -HLK HLKV L // #L #W #HKL #HKVL #HWV + elim (IHV12 … HKVL … HWV ?) -HKVL HWV /2/ -Hdetd #W1 #HW1 #HWV1 + elim (lift_trans_le … HWV1 … HV12 ?) -HWV1 HV12 V1 // >arith_a2 /3 width=6/ +| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd + elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct -X; + elim (IHV12 … HLK … HWV1 ?) -IHV12 // + elim (IHU12 … HTU1 ?) -IHU12 HTU1 [3: /2/ |4: @drop_skip // |2: skip ] -HLK HWV1 Hdetd /3 width=5/ (**) (* just /3 width=5/ is too slow *) +| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd + elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct -X; + elim (IHV12 … HLK … HWV1 ?) -IHV12 HWV1 // + elim (IHU12 … HLK … HTU1 ?) -IHU12 HLK HTU1 // /3 width=5/ +] +qed. + +lemma tps_inv_lift1_ge: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 → + ∀K,d,e. ↓[d, e] L ≡ K → ∀T1. ↑[d, e] T1 ≡ U1 → + d + e ≤ dt → + ∃∃T2. K ⊢ T1 [dt - e, et] ≫ T2 & ↑[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #H elim H -H L U1 U2 dt et +[ #L #k #dt #et #K #d #e #_ #T1 #H #_ + lapply (lift_inv_sort2 … H) -H #H destruct -T1 /2/ +| #L #i #dt #et #K #d #e #_ #T1 #H #_ + elim (lift_inv_lref2 … H) -H * #Hid #H destruct -T1 /3/ +| #L #KV #V #V1 #V2 #i #dt #et #Hdti #Hidet #HLKV #HV1 #HV12 #_ #K #d #e #HLK #T1 #H #Hdedt + lapply (transitive_le … Hdedt … Hdti) #Hdei + lapply (plus_le_weak … Hdedt) -Hdedt #Hedt + lapply (plus_le_weak … Hdei) #Hei + <(arith_h1 ? ? ? e ? ?) in HV1 // #HV1 + lapply (lift_inv_lref2_ge … H … Hdei) -H #H destruct -T1; + lapply (drop_conf_ge … HLK … HLKV ?) -HLK HLKV L // #HKV + elim (lift_split … HV12 d (i - e + 1) ? ? ?) -HV12; [2,3,4: normalize /2/ ] -Hdei >arith_e2 // #V0 #HV10 #HV02 + @ex2_1_intro + [2: @tps_subst [4: /2/ |6,7,8: // |1,2,3: skip |5: @arith5 // ] + |1: skip + | // + ] (**) (* explicitc constructors *) +| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd + elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct -X; + lapply (plus_le_weak … Hdetd) #Hedt + elim (IHV12 … HLK … HWV1 ?) -IHV12 // #W2 #HW12 #HWV2 + elim (IHU12 … HTU1 ?) -IHU12 HTU1 [4: @drop_skip // |2: skip |3: /2/ ] + IHV12 // >IHT12 // +| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX + elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #H destruct -X + >IHV12 // >IHT12 // +] +qed. +(* + Theorem subst0_gen_lift_ge : (u,t1,x:?; i,h,d:?) (subst0 i u (lift h d t1) x) -> + (le (plus d h) i) -> + (EX t2 | x = (lift h d t2) & (subst0 (minus i h) u t1 t2)). + + Theorem subst0_gen_lift_rev_ge: (t1,v,u2:?; i,h,d:?) + (subst0 i v t1 (lift h d u2)) -> + (le (plus d h) i) -> + (EX u1 | (subst0 (minus i h) v u1 u2) & + t1 = (lift h d u1) + ). + + + Theorem subst0_gen_lift_rev_lelt: (t1,v,u2:?; i,h,d:?) + (subst0 i v t1 (lift h d u2)) -> + (le d i) -> (lt i (plus d h)) -> + (EX u1 | t1 = (lift (minus (plus d h) (S i)) (S i) u1)). +*) diff --git a/matita/matita/lib/lambda-delta/substitution/tps_split.ma b/matita/matita/lib/lambda-delta/substitution/tps_split.ma new file mode 100644 index 000000000..23cdb39df --- /dev/null +++ b/matita/matita/lib/lambda-delta/substitution/tps_split.ma @@ -0,0 +1,58 @@ +(* + ||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/tps_lift.ma". + +(* PARTIAL SUBSTITUTION ON TERMS ********************************************) + +(* Split properties *********************************************************) + +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/ +| /2/ +| #L #K #V #V1 #V2 #i #d #e #Hdi #Hide #HLK #HV1 #HV12 #IHV12 #j #Hdj #Hjde + elim (lt_or_ge i j) #Hij + [ -HV1 Hide; + lapply (drop_fwd_drop2 … HLK) #HLK' + elim (IHV12 (j - i - 1) ? ?) -IHV12; normalize /2/ -Hjde arith_b2 // #W1 #HVW1 #HWV1 + generalize in match HVW1 generalize in match Hij -HVW1 (**) (* rewriting in the premises, rewrites in the goal too *) + >(plus_minus_m_m_comm … Hdj) in ⊢ (% → % → ?) -Hdj #Hij' #HVW1 + elim (lift_total W1 0 (i + 1)) #W2 #HW12 + lapply (tps_lift_ge … HWV1 … HLK' HW12 HV12 ?) -HWV1 HLK' HV12 // >arith_a2 /3 width=6/ + | -IHV12 Hdi Hdj; + generalize in match HV1 generalize in match Hide -HV1 Hide (**) (* rewriting in the premises, rewrites in the goal too *) + >(plus_minus_m_m_comm … Hjde) in ⊢ (% → % → ?) -Hjde #Hide #HV1 + @ex2_1_intro [2: @tps_lref |1: skip | /2 width=6/ ] (**) (* /3 width=6 is too slow *) + ] +| #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 + @ex2_1_intro [2,3: @tps_bind | skip ] (**) (* explicit constructors *) + [3: @HV1 |4: @HT1 |5: // |1,2: skip | /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. + +lemma tps_inv_lift1_up: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 → + ∀K,d,e. ↓[d, e] L ≡ K → ∀T1. ↑[d, e] T1 ≡ U1 → + d ≤ dt → dt ≤ d + e → d + e ≤ dt + et → + ∃∃T2. K ⊢ T1 [d, dt + et - (d + e)] ≫ T2 & ↑[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #HU12 #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet +elim (tps_split_up … HU12 (d + e) ? ?) -HU12 // -Hdedet #U #HU1 #HU2 +lapply (tps_weak … HU1 d e ? ?) -HU1 // + (u1,u:?; i:?) (subst0 i u u1 u2) -> + (EX t | (subst0 j u1 t1 t) & (subst0 (S (plus i j)) u t t2)). + + Theorem subst0_subst0_back: (t1,t2,u2:?; j:?) (subst0 j u2 t1 t2) -> + (u1,u:?; i:?) (subst0 i u u2 u1) -> + (EX t | (subst0 j u1 t1 t) & (subst0 (S (plus i j)) u t2 t)). + +*) diff --git a/matita/matita/lib/lambda-delta/xoa.conf.xml b/matita/matita/lib/lambda-delta/xoa.conf.xml index 3330813ae..d25c5dcca 100644 --- a/matita/matita/lib/lambda-delta/xoa.conf.xml +++ b/matita/matita/lib/lambda-delta/xoa.conf.xml @@ -2,7 +2,7 @@
$(MATITA_RT_BASE_DIR) -
diff --git a/matita/matita/lib/lambda-delta/xoa.ma b/matita/matita/lib/lambda-delta/xoa.ma index 582b2d25f..68a0776dc 100644 --- a/matita/matita/lib/lambda-delta/xoa.ma +++ b/matita/matita/lib/lambda-delta/xoa.ma @@ -125,11 +125,3 @@ inductive or4 (P0,P1,P2,P3:Prop) : Prop ≝ interpretation "multiple disjunction connective (4)" 'Or P0 P1 P2 P3 = (or4 P0 P1 P2 P3). -(* multiple conjunction connective (3) *) - -inductive and3 (P0,P1,P2:Prop) : Prop ≝ - | and3_intro: P0 → P1 → P2 → and3 ? ? ? -. - -interpretation "multiple conjunction connective (3)" 'And P0 P1 P2 = (and3 P0 P1 P2). - diff --git a/matita/matita/lib/lambda-delta/xoa_notation.ma b/matita/matita/lib/lambda-delta/xoa_notation.ma index ce9460eb6..b8270b660 100644 --- a/matita/matita/lib/lambda-delta/xoa_notation.ma +++ b/matita/matita/lib/lambda-delta/xoa_notation.ma @@ -136,9 +136,3 @@ notation "hvbox(∨∨ term 29 P0 break | term 29 P1 break | term 29 P2 break | non associative with precedence 30 for @{ 'Or $P0 $P1 $P2 $P3 }. -(* multiple conjunction connective (3) *) - -notation "hvbox(∧∧ term 34 P0 break & term 34 P1 break & term 34 P2)" - non associative with precedence 35 - for @{ 'And $P0 $P1 $P2 }. -