From 093f9476ddf96034b89e6ad443f74bcc6c067912 Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Tue, 23 Aug 2011 17:07:23 +0000 Subject: [PATCH] - confluence of parallel substitution (tps) closed! (a bug in an inversion lemma was in the way) - some refactoring --- .../lambda-delta/Basic-2/substitution/tps.ma | 43 ++++++++--- .../Basic-2/substitution/tps_lift.ma | 11 +++ .../Basic-2/substitution/tps_split.ma | 55 -------------- .../Basic-2/substitution/tps_tps.ma | 71 ++++++++----------- .../lambda-delta/Ground-2/xoa.conf.xml | 1 + .../contribs/lambda-delta/Ground-2/xoa.ma | 8 +++ .../lambda-delta/Ground-2/xoa_notation.ma | 10 +++ 7 files changed, 94 insertions(+), 105 deletions(-) delete mode 100644 matita/matita/contribs/lambda-delta/Basic-2/substitution/tps_split.ma diff --git a/matita/matita/contribs/lambda-delta/Basic-2/substitution/tps.ma b/matita/matita/contribs/lambda-delta/Basic-2/substitution/tps.ma index bc82e4e6d..fc0851097 100644 --- a/matita/matita/contribs/lambda-delta/Basic-2/substitution/tps.ma +++ b/matita/matita/contribs/lambda-delta/Basic-2/substitution/tps.ma @@ -86,27 +86,52 @@ lapply (tps_weak … HT12 0 (d + e) ? ?) -HT12 // #HT12 lapply (tps_weak_top … HT12) // 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/ +| /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 + @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. + (* Basic inversion lemmas ***************************************************) lemma tps_inv_lref1_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. + T2 = #i ∨ + ∃∃K,V. d ≤ i & i < d + e & + ↓[O, i] L ≡ K. 𝕓{Abbr} V & + ↑[O, i + 1] V ≡ T2. #L #T1 #T2 #d #e * -L T1 T2 d e [ #L #k #d #e #i #H destruct | /2/ -| #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #j #H destruct -i /3 width=7/ +| #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #j #H destruct -i /3/ | #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct | #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct ] qed. lemma tps_inv_lref1: ∀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. + T2 = #i ∨ + ∃∃K,V. d ≤ i & i < d + e & + ↓[O, i] L ≡ K. 𝕓{Abbr} V & + ↑[O, i + 1] V ≡ T2. /2/ qed. lemma tps_inv_bind1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ≫ U2 → diff --git a/matita/matita/contribs/lambda-delta/Basic-2/substitution/tps_lift.ma b/matita/matita/contribs/lambda-delta/Basic-2/substitution/tps_lift.ma index e5ad25ff1..d15b59535 100644 --- a/matita/matita/contribs/lambda-delta/Basic-2/substitution/tps_lift.ma +++ b/matita/matita/contribs/lambda-delta/Basic-2/substitution/tps_lift.ma @@ -170,3 +170,14 @@ qed. (le d i) -> (lt i (plus d h)) -> (EX u1 | t1 = (lift (minus (plus d h) (S i)) (S i) u1)). *) + +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 // (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 - @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 // (lift_mono … HVT1 … HVT2) -HVT1 /2/ - | @ex2_1_comm @tps_conf_subst_subst_lt /width=15/ - ] + | * #K2 #V2 #_ #_ #HLK2 #HVT2 + lapply (drop_mono … HLK1 … HLK2) -HLK1 #H destruct -V1 K1 + >(lift_mono … HVT1 … HVT2) -HVT1 /2/ ] | #L #I #V0 #V1 #T0 #T1 #d #e #_ #_ #IHV01 #IHT01 #X #HX - elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct -X; + elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct -X; elim (IHV01 … HV02) -IHV01 HV02 #V #HV1 #HV2 - elim (IHT01 … HT02) -IHT01 HT02 #T #HT1 #HT2 + elim (IHT01 … HT02) -IHT01 HT02 #T #HT1 #HT2 @ex2_1_intro [2: @tps_bind [4: @(tps_leq_repl … HT1) /2/ |2: skip ] |1: skip |3: @tps_bind [2: @(tps_leq_repl … HT2) /2/ ] ] // (**) (* /5/ is too slow *) | #L #I #V0 #V1 #T0 #T1 #d #e #_ #_ #IHV01 #IHT01 #X #HX - elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct -X; + elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct -X; elim (IHV01 … HV02) -IHV01 HV02; elim (IHT01 … HT02) -IHT01 HT02 /3 width=5/ ] qed. +theorem tps_trans_down: ∀L,T1,T0,d1,e1. L ⊢ T1 [d1, e1] ≫ T0 → + ∀T2,d2,e2. L ⊢ T0 [d2, e2] ≫ T2 → d2 + e2 ≤ d1 → + ∃∃T. L ⊢ T1 [d2, e2] ≫ T & L ⊢ T [d1, e1] ≫ T2. +#L #T1 #T0 #d1 #e1 #H elim H -L T1 T0 d1 e1 +[ /2/ +| /2/ +| #L #K #V #W #i1 #d1 #e1 #Hdi1 #Hide1 #HLK #HVW #T2 #d2 #e2 #HWT2 #Hde2d1 + lapply (transitive_le … Hde2d1 Hdi1) -Hde2d1 #Hde2i1 + lapply (tps_weak … HWT2 0 (i1 + 1) ? ?) -HWT2; normalize /2/ -Hde2i1 #HWT2 + <(tps_inv_lift1_eq … HWT2 … HVW) -HWT2 /4/ +| #L #I #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1 + elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct -X; + lapply (tps_leq_repl … HT02 (L. 𝕓{I} V1) ?) -HT02 /2/ #HT02 + elim (IHV10 … HV02 ?) -IHV10 HV02 // #V + elim (IHT10 … HT02 ?) -IHT10 HT02 [2: /2/ ] #T #HT1 #HT2 + lapply (tps_leq_repl … HT2 (L. 𝕓{I} V) ?) -HT2 /3 width=6/ +| #L #I #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1 + elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct -X; + elim (IHV10 … HV02 ?) -IHV10 HV02 // + elim (IHT10 … HT02 ?) -IHT10 HT02 // /3 width=6/ +] +qed. (* Theorem subst0_subst0: (t1,t2,u2:?; j:?) (subst0 j u2 t1 t2) -> (u1,u:?; i:?) (subst0 i u u1 u2) -> diff --git a/matita/matita/contribs/lambda-delta/Ground-2/xoa.conf.xml b/matita/matita/contribs/lambda-delta/Ground-2/xoa.conf.xml index f0ba4ec81..e7f9feef4 100644 --- a/matita/matita/contribs/lambda-delta/Ground-2/xoa.conf.xml +++ b/matita/matita/contribs/lambda-delta/Ground-2/xoa.conf.xml @@ -19,6 +19,7 @@ 3 1 3 2 3 3 + 4 2 4 3 4 4 5 3 diff --git a/matita/matita/contribs/lambda-delta/Ground-2/xoa.ma b/matita/matita/contribs/lambda-delta/Ground-2/xoa.ma index 68a0776dc..a9deb0fa1 100644 --- a/matita/matita/contribs/lambda-delta/Ground-2/xoa.ma +++ b/matita/matita/contribs/lambda-delta/Ground-2/xoa.ma @@ -56,6 +56,14 @@ inductive ex3_3 (A0,A1,A2:Type[0]) (P0,P1,P2:A0→A1→A2→Prop) : Prop ≝ interpretation "multiple existental quantifier (3, 3)" 'Ex P0 P1 P2 = (ex3_3 ? ? ? P0 P1 P2). +(* multiple existental quantifier (4, 2) *) + +inductive ex4_2 (A0,A1:Type[0]) (P0,P1,P2,P3:A0→A1→Prop) : Prop ≝ + | ex4_2_intro: ∀x0,x1. P0 x0 x1 → P1 x0 x1 → P2 x0 x1 → P3 x0 x1 → ex4_2 ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (4, 2)" 'Ex P0 P1 P2 P3 = (ex4_2 ? ? P0 P1 P2 P3). + (* multiple existental quantifier (4, 3) *) inductive ex4_3 (A0,A1,A2:Type[0]) (P0,P1,P2,P3:A0→A1→A2→Prop) : Prop ≝ diff --git a/matita/matita/contribs/lambda-delta/Ground-2/xoa_notation.ma b/matita/matita/contribs/lambda-delta/Ground-2/xoa_notation.ma index b8270b660..5f747314c 100644 --- a/matita/matita/contribs/lambda-delta/Ground-2/xoa_notation.ma +++ b/matita/matita/contribs/lambda-delta/Ground-2/xoa_notation.ma @@ -64,6 +64,16 @@ notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break non associative with precedence 20 for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P2) }. +(* multiple existental quantifier (4, 2) *) + +notation > "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.λ${ident x1}.$P0) (λ${ident x0}.λ${ident x1}.$P1) (λ${ident x0}.λ${ident x1}.$P2) (λ${ident x0}.λ${ident x1}.$P3) }. + +notation < "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P3) }. + (* multiple existental quantifier (4, 3) *) notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)" -- 2.39.2